Skip to content
Snippets Groups Projects
Commit fd9a4c7a authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Remove ◇ in def of ⊑.

parent 7a449ae5
No related branches found
No related tags found
No related merge requests found
Pipeline #34182 passed
...@@ -259,8 +259,7 @@ Section channel. ...@@ -259,8 +259,7 @@ Section channel.
iIntros "_". wp_pures. iApply "HΦ". iLeft. iSplit; [done|]. iIntros "_". wp_pures. iApply "HΦ". iLeft. iSplit; [done|].
iExists γ, Left, l, r, lk. eauto 10 with iFrame. } iExists γ, Left, l, r, lk. eauto 10 with iFrame. }
wp_apply (lpop_spec with "Hr"); iIntros (v') "[% Hr]"; simplify_eq/=. wp_apply (lpop_spec with "Hr"); iIntros (v') "[% Hr]"; simplify_eq/=.
iMod (iProto_recv_l with "Hctx H") as "H". wp_pures. iMod (iProto_recv_l with "Hctx H") as (q) "(Hctx & H & Hm)". wp_pures.
iMod "H" as (q) "(Hctx & H & Hm)".
rewrite iMsg_base_eq. iDestruct (iMsg_texist with "Hm") as (x <-) "[Hp HP]". rewrite iMsg_base_eq. iDestruct (iMsg_texist with "Hm") as (x <-) "[Hp HP]".
wp_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. wp_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|].
iIntros "_". wp_pures. iApply "HΦ". iRight. iExists x. iSplit; [done|]. iIntros "_". wp_pures. iApply "HΦ". iRight. iExists x. iSplit; [done|].
...@@ -272,8 +271,7 @@ Section channel. ...@@ -272,8 +271,7 @@ Section channel.
iIntros "_". wp_pures. iApply "HΦ". iLeft. iSplit; [done|]. iIntros "_". wp_pures. iApply "HΦ". iLeft. iSplit; [done|].
iExists γ, Right, l, r, lk. eauto 10 with iFrame. } iExists γ, Right, l, r, lk. eauto 10 with iFrame. }
wp_apply (lpop_spec with "Hl"); iIntros (v') "[% Hl]"; simplify_eq/=. wp_apply (lpop_spec with "Hl"); iIntros (v') "[% Hl]"; simplify_eq/=.
iMod (iProto_recv_r with "Hctx H") as "H". wp_pures. iMod (iProto_recv_r with "Hctx H") as (q) "(Hctx & H & Hm)". wp_pures.
iMod "H" as (q) "(Hctx & H & Hm)".
rewrite iMsg_base_eq. iDestruct (iMsg_texist with "Hm") as (x <-) "[Hp HP]". rewrite iMsg_base_eq. iDestruct (iMsg_texist with "Hm") as (x <-) "[Hp HP]".
wp_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. wp_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|].
iIntros "_". wp_pures. iApply "HΦ". iRight. iExists x. iSplit; [done|]. iIntros "_". wp_pures. iApply "HΦ". iRight. iExists x. iSplit; [done|].
......
This diff is collapsed.
...@@ -150,7 +150,7 @@ Section choice_example. ...@@ -150,7 +150,7 @@ Section choice_example.
(** Weakening of select *) (** Weakening of select *)
iApply (lty_le_trans _ prot1). iApply (lty_le_trans _ prot1).
{ {
iApply lty_le_branch. iIntros "!>". iApply lty_le_branch.
rewrite big_sepM2_insert=> //. rewrite big_sepM2_insert=> //.
iSplit. iSplit.
- iApply lty_le_recv; [iApply lty_le_refl | ]. - iApply lty_le_recv; [iApply lty_le_refl | ].
...@@ -163,7 +163,7 @@ Section choice_example. ...@@ -163,7 +163,7 @@ Section choice_example.
(** Swap recv/select *) (** Swap recv/select *)
iApply (lty_le_trans _ prot2). iApply (lty_le_trans _ prot2).
{ {
iApply lty_le_branch. iIntros "!>". iApply lty_le_branch.
rewrite big_sepM2_insert=> //. rewrite big_sepM2_insert=> //.
iSplit. iSplit.
- iApply lty_le_swap_recv_select. - iApply lty_le_swap_recv_select.
...@@ -223,7 +223,7 @@ Section choice_example. ...@@ -223,7 +223,7 @@ Section choice_example.
(** Swap recv/send *) (** Swap recv/send *)
iApply (lty_le_trans _ prot4). iApply (lty_le_trans _ prot4).
{ {
iApply lty_le_select. iIntros "!>". iApply lty_le_select.
rewrite big_sepM2_insert=> //. iSplit. rewrite big_sepM2_insert=> //. iSplit.
- iApply lty_le_branch. iIntros "!>". - iApply lty_le_branch. iIntros "!>".
rewrite big_sepM2_insert=> //. iSplit. rewrite big_sepM2_insert=> //. iSplit.
...@@ -239,7 +239,7 @@ Section choice_example. ...@@ -239,7 +239,7 @@ Section choice_example.
} }
iApply (lty_le_trans _ prot5). iApply (lty_le_trans _ prot5).
{ {
iApply lty_le_select. iIntros "!>". iApply lty_le_select.
rewrite big_sepM2_insert=> //. iSplit. rewrite big_sepM2_insert=> //. iSplit.
- iApply lty_le_branch. iIntros "!>". - iApply lty_le_branch. iIntros "!>".
rewrite big_sepM2_insert=> //. iSplit. rewrite big_sepM2_insert=> //. iSplit.
...@@ -257,7 +257,7 @@ Section choice_example. ...@@ -257,7 +257,7 @@ Section choice_example.
(** Swap branch/send *) (** Swap branch/send *)
iApply (lty_le_trans _ prot6). iApply (lty_le_trans _ prot6).
{ {
iApply lty_le_select. iIntros "!>". iApply lty_le_select.
rewrite big_sepM2_insert=> //. iSplit. rewrite big_sepM2_insert=> //. iSplit.
- iApply (lty_le_swap_branch_send _ - iApply (lty_le_swap_branch_send _
(<[1%Z:=(<??> TY Sr; Tr)%lty]> (<[2%Z:=(<??> TY Ss; Ts)%lty]> ))). (<[1%Z:=(<??> TY Sr; Tr)%lty]> (<[2%Z:=(<??> TY Ss; Ts)%lty]> ))).
...@@ -266,7 +266,7 @@ Section choice_example. ...@@ -266,7 +266,7 @@ Section choice_example.
((<[1%Z:=(<??> TY Sr'; Tr')%lty]> (<[2%Z:=(<??> TY Ss; Ts')%lty]> )))). ((<[1%Z:=(<??> TY Sr'; Tr')%lty]> (<[2%Z:=(<??> TY Ss; Ts')%lty]> )))).
} }
(** Weaken branch *) (** Weaken branch *)
iApply lty_le_select. iIntros "!>". iApply lty_le_select.
rewrite big_sepM2_insert=> //. iSplit=> //. rewrite big_sepM2_insert=> //. iSplit=> //.
- iApply lty_le_send; [iApply lty_le_refl|]. - iApply lty_le_send; [iApply lty_le_refl|].
iApply lty_le_branch_subseteq. iApply lty_le_branch_subseteq.
......
...@@ -371,15 +371,15 @@ Section subtyping_rules. ...@@ -371,15 +371,15 @@ Section subtyping_rules.
Qed. Qed.
Lemma lty_le_select (Ss1 Ss2 : gmap Z (lsty Σ)) : Lemma lty_le_select (Ss1 Ss2 : gmap Z (lsty Σ)) :
([ map] S1;S2 Ss1; Ss2, S1 <: S2) -∗ ([ map] S1;S2 Ss1; Ss2, (S1 <: S2)) -∗
lty_select Ss1 <: lty_select Ss2. lty_select Ss1 <: lty_select Ss2.
Proof. Proof.
iIntros "#H !>" (x); iDestruct 1 as %[S2 HSs2]. iExists x. iIntros "#H !>" (x); iDestruct 1 as %[S2 HSs2]. iExists x.
iDestruct (big_sepM2_forall with "H") as "{H} [>% H]". iDestruct (big_sepM2_forall with "H") as (?) "{H} H".
assert (is_Some (Ss1 !! x)) as [S1 HSs1] by naive_solver. assert (is_Some (Ss1 !! x)) as [S1 HSs1] by naive_solver.
iSpecialize ("H" with "[//] [//]").
rewrite HSs1. iSplitR; [by eauto|]. rewrite HSs1. iSplitR; [by eauto|].
iIntros "!>". rewrite !lookup_total_alt HSs1 HSs2 /=. iIntros "!>". by rewrite !lookup_total_alt HSs1 HSs2 /=.
by iApply ("H" with "[] []").
Qed. Qed.
Lemma lty_le_select_subseteq (Ss1 Ss2 : gmap Z (lsty Σ)) : Lemma lty_le_select_subseteq (Ss1 Ss2 : gmap Z (lsty Σ)) :
Ss2 Ss1 Ss2 Ss1
...@@ -392,15 +392,15 @@ Section subtyping_rules. ...@@ -392,15 +392,15 @@ Section subtyping_rules.
Qed. Qed.
Lemma lty_le_branch (Ss1 Ss2 : gmap Z (lsty Σ)) : Lemma lty_le_branch (Ss1 Ss2 : gmap Z (lsty Σ)) :
([ map] S1;S2 Ss1; Ss2, S1 <: S2) -∗ ([ map] S1;S2 Ss1; Ss2, (S1 <: S2)) -∗
lty_branch Ss1 <: lty_branch Ss2. lty_branch Ss1 <: lty_branch Ss2.
Proof. Proof.
iIntros "#H !>" (x); iDestruct 1 as %[S1 HSs1]. iExists x. iIntros "#H !>" (x); iDestruct 1 as %[S1 HSs1]. iExists x.
iDestruct (big_sepM2_forall with "H") as "{H} [>% H]". iDestruct (big_sepM2_forall with "H") as (?) "{H} H".
assert (is_Some (Ss2 !! x)) as [S2 HSs2] by naive_solver. assert (is_Some (Ss2 !! x)) as [S2 HSs2] by naive_solver.
iSpecialize ("H" with "[//] [//]").
rewrite HSs2. iSplitR; [by eauto|]. rewrite HSs2. iSplitR; [by eauto|].
iIntros "!>". rewrite !lookup_total_alt HSs1 HSs2 /=. iIntros "!>". by rewrite !lookup_total_alt HSs1 HSs2 /=.
by iApply ("H" with "[] []").
Qed. Qed.
Lemma lty_le_branch_subseteq (Ss1 Ss2 : gmap Z (lsty Σ)) : Lemma lty_le_branch_subseteq (Ss1 Ss2 : gmap Z (lsty Σ)) :
Ss1 Ss2 Ss1 Ss2
......
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