Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
L
lambda-rust
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Wiki
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
Build
Pipelines
Jobs
Pipeline schedules
Artifacts
Deploy
Releases
Model registry
Operate
Environments
Monitor
Service Desk
Help
Help
Support
GitLab documentation
Compare GitLab plans
Community forum
Contribute to GitLab
Provide feedback
Terms and privacy
Keyboard shortcuts
?
Snippets
Groups
Projects
Show more breadcrumbs
Iris
lambda-rust
Commits
e189f4ad
Commit
e189f4ad
authored
3 years ago
by
Hai Dang
Browse files
Options
Downloads
Patches
Plain Diff
Use Atomic ptsto and cancelable inv for spawn and join
parent
e83da50a
No related branches found
Branches containing commit
No related tags found
No related merge requests found
Pipeline
#62189
passed
3 years ago
Stage: build
Changes
3
Pipelines
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
theories/lang/spawn.v
+155
-105
155 additions, 105 deletions
theories/lang/spawn.v
theories/typing/lib/join.v
+2
-2
2 additions, 2 deletions
theories/typing/lib/join.v
theories/typing/lib/spawn.v
+5
-5
5 additions, 5 deletions
theories/typing/lib/spawn.v
with
162 additions
and
112 deletions
theories/lang/spawn.v
+
155
−
105
View file @
e189f4ad
From
iris
.
base_logic
.
lib
Require
Import
invariants
.
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
algebra
Require
Import
excl
.
From
lrust
.
lang
Require
Import
notation
.
From
lrust
.
lang
Require
Import
notation
.
From
gpfsl
.
logic
Require
Import
view_invariants
.
From
gpfsl
.
logic
Require
Import
proofmode
atomics
view_invariants
.
From
gpfsl
.
gps
Require
Import
surface
protocols
escrows
.
From
iris
.
prelude
Require
Import
options
.
From
iris
.
prelude
Require
Import
options
.
...
@@ -16,8 +14,8 @@ Definition spawn : val :=
...
@@ -16,8 +14,8 @@ Definition spawn : val :=
Definition
finish
:
val
:=
Definition
finish
:
val
:=
λ
:
[
"c"
;
"v"
],
λ
:
[
"c"
;
"v"
],
"c"
+
ₗ
#
1
<-
"v"
;;
(* Store data (non-atomically). *)
"c"
+
ₗ
#
1
<-
"v"
;;
(* Store data (non-atomically). *)
"c"
<-
ʳᵉˡ
#
1
;;
(* Signal that we finished (atomically!) *)
"c"
<-
ʳᵉˡ
#
1
(* Signal that we finished (atomically!) *)
#☠
.
.
Definition
join
:
val
:=
Definition
join
:
val
:=
rec
:
"join"
[
"c"
]
:=
rec
:
"join"
[
"c"
]
:=
if
:
!
ᵃᶜ
"c"
then
if
:
!
ᵃᶜ
"c"
then
...
@@ -30,148 +28,200 @@ Definition join : val :=
...
@@ -30,148 +28,200 @@ Definition join : val :=
(** The CMRA & functor we need. *)
(** The CMRA & functor we need. *)
Class
spawnG
Σ
:=
SpawnG
{
Class
spawnG
Σ
:=
SpawnG
{
spawn_tokG
:>
inG
Σ
(
exclR
unitO
);
spawn_gpsG
:>
gpsG
Σ
unitProtocol
;
spawn_atomG
:>
atomicG
Σ
;
spawn_atomG
:>
atomicG
Σ
;
spawn_viewG
:>
view_invG
Σ
}
.
spawn_viewG
:>
view_invG
Σ
}
.
Definition
spawnΣ
:
gFunctors
:=
Definition
spawnΣ
:
gFunctors
:=
#
[
atomicΣ
;
view_invΣ
]
.
#
[
GFunctor
(
exclR
unitO
);
gpsΣ
unitProtocol
;
atomicΣ
;
view_invΣ
]
.
Global
Instance
subG_spawnΣ
{
Σ
}
:
subG
spawnΣ
Σ
→
spawnG
Σ
.
Global
Instance
subG_spawnΣ
{
Σ
}
:
subG
spawnΣ
Σ
→
spawnG
Σ
.
Proof
.
solve_inG
.
Qed
.
Proof
.
solve_inG
.
Qed
.
(** Now we come to the Iris part of the proof. *)
Section
proof
.
Section
proof
.
Context
`{
!
noprolG
Σ
,
!
spawnG
Σ
,
!
view_invG
Σ
}
(
N
:
namespace
)
.
Context
`{
!
noprolG
Σ
,
!
spawnG
Σ
}
(
N
:
namespace
)
.
Local
Notation
vProp
:=
(
vProp
Σ
)
.
Local
Notation
vProp
:=
(
vProp
Σ
)
.
Local
Notation
tok
γ
:=
(
own
γ
(
Excl
()))
.
Implicit
Types
(
c
:
loc
)
(
Ψ
:
val
→
vProp
)
(
v
:
val
)
.
Implicit
Types
(
c
:
loc
)
(
Ψ
:
val
→
vProp
)
(
v
:
val
)
.
Definition
finishEscrow
γc
γe
γi
c
Ψ
:=
Local
Definition
mp_inv_reclaim'_def
γc
γi
c
Ψ
:
vProp
:=
([
es
⎡
tok
γe
⎤
⇝
((
∃
t
,
GPS_SWWriter
c
t
()
#
1
γc
)
∗
view_tok
γi
(
1
/
2
)
%
Qp
∗
(
∃
ζ
(
b
:
bool
)
t0
V0
,
∃
v
,
(
c
>>
1
)
↦
v
∗
Ψ
v
)])
%
I
.
c
sw
↦
{
γc
}
ζ
∗
let
ζ0
:
absHist
:=
{[
t0
:=
(
#
0
,
V0
)]}
in
(* GPS protocol interpretation *)
match
b
with
Definition
spawn_interp
γe
γi
Ψ
:
interpO
Σ
unitProtocol
:=
|
false
=>
⌜
ζ
=
ζ0
⌝
(
λ
_
c
γc
_
_
v
,
∃
b
:
bool
,
⌜
v
=
#
b
⌝
∗
|
true
=>
∃
t1
V1
,
⌜
(
t0
<
t1
)
%
positive
∧
ζ
=
<
[
t1
:=
(
#
1
,
V1
)]
>
ζ0
⌝
∗
if
b
then
finishEscrow
γc
γe
γi
c
Ψ
else
True
)
%
I
.
@
{
V1
}
(
∃
v
,
(
c
>>
1
)
↦
v
∗
Ψ
v
∗
view_tok
γi
(
1
/
2
))
end
Let
IW
:
loc
→
gname
→
time
→
unitProtocol
→
val
→
vProp
:=
)
%
I
.
(
λ
_
_
_
_
v
,
⌜
v
=
#
false
⌝
)
%
I
.
Local
Definition
mp_inv_reclaim'_aux
:
seal
(
@
mp_inv_reclaim'_def
)
.
Proof
.
by
eexists
.
Qed
.
Local
Definition
mp_inv_reclaim'
:=
unseal
(
@
mp_inv_reclaim'_aux
)
.
Definition
finish_handle
γc
γe
c
Ψ
:
vProp
:=
Local
Definition
mp_inv_reclaim'_eq
:
@
mp_inv_reclaim'
=
_
:=
seal_eq
_
.
(
∃
γi
t
v
,
(
c
>>
1
)
↦
v
∗
GPS_vSP_Writer
N
(
spawn_interp
γe
γi
Ψ
)
IW
c
t
()
#
0
γc
γi
∗
(* We don't care about objectivity when using view_inv *)
Local
Definition
mp_inv_reclaim
γc
γi
c
Ψ
:=
view_inv
γi
N
(
mp_inv_reclaim'
γc
γi
c
Ψ
)
.
Definition
finish_handle
γc
c
Ψ
:
vProp
:=
(
∃
γi
t
V
v
,
(
c
>>
1
)
↦
v
∗
c
sw
⊒
{
γc
}
{[
t
:=
(
#
0
,
V
)]}
∗
mp_inv_reclaim
γc
γi
c
Ψ
∗
view_tok
γi
(
1
/
2
)
%
Qp
)
%
I
.
view_tok
γi
(
1
/
2
)
%
Qp
)
%
I
.
Definition
join_handle
γc
γe
c
Ψ
:
vProp
:=
Definition
join_handle
γc
c
Ψ
:
vProp
:=
(
∃
γi
(
Ψ'
:
val
→
vProp
)
t
,
(
∃
γi
(
Ψ'
:
val
→
vProp
)
t
V
,
GPS_vSP_Reader
N
(
spawn_interp
γe
γi
Ψ'
)
IW
c
t
()
#
0
γc
γi
c
sn
⊒
{
γc
}
{[
t
:=
(
#
0
,
V
)]}
∗
⎡
†
c
…
2
∗
tok
γe
⎤
∗
mp_inv_reclaim
γc
γi
c
Ψ'
∗
⎡
†
c
…
2
⎤
∗
view_tok
γi
(
1
/
2
)
%
Qp
∗
view_tok
γi
(
1
/
2
)
%
Qp
∗
□
(
∀
v
,
Ψ'
v
-∗
Ψ
v
))
%
I
.
∗
□
(
∀
v
,
Ψ'
v
-∗
Ψ
v
))
%
I
.
Global
Instance
finish_handle_contractive
n
γc
γe
c
:
Global
Instance
finish_handle_contractive
n
γc
c
:
Proper
(
pointwise_relation
val
(
dist_later
n
)
==>
dist
n
)
(
finish_handle
γc
γe
c
)
.
Proper
(
pointwise_relation
val
(
dist_later
n
)
==>
dist
n
)
(
finish_handle
γc
c
)
.
Proof
.
Proof
.
move
=>
???
.
do
3
(
apply
bi
.
exist_ne
=>
?)
.
rewrite
/
finish_handle
/
mp_inv_reclaim
mp_inv_reclaim'_eq
/
mp_inv_reclaim'_def
.
apply
bi
.
sep_ne
;
[
done
|]
.
apply
bi
.
sep_ne
;
[|
done
]
.
solve_contractive
.
apply
GPS_vSP_Writer_contractive
;
[|
done
]
.
destruct
n
;
[
done
|]
.
simpl
=>
??????
.
apply
bi
.
exist_ne
=>
b'
.
apply
bi
.
sep_ne
;
[
done
|]
.
destruct
b'
;
[|
done
]
.
solve_proper
.
Qed
.
Qed
.
Global
Instance
join_handle_ne
n
γc
γe
c
:
Global
Instance
join_handle_ne
n
γc
c
:
Proper
(
pointwise_relation
val
(
dist
n
)
==>
dist
n
)
(
join_handle
γc
γe
c
)
.
Proper
(
pointwise_relation
val
(
dist
n
)
==>
dist
n
)
(
join_handle
γc
c
)
.
Proof
.
solve_proper
.
Qed
.
Proof
.
solve_proper
.
Qed
.
(** The main proofs. *)
(** The main proofs. *)
Lemma
spawn_spec
Ψ
e
(
f
:
val
)
tid
:
Lemma
spawn_spec
Ψ
e
(
f
:
val
)
tid
:
IntoVal
e
f
→
IntoVal
e
f
→
{{{
∀
γc
γe
c
tid'
,
finish_handle
γc
γe
c
Ψ
-∗
WP
f
[
#
c
]
@
tid'
;
⊤
{{
_,
True
}}
}}}
{{{
∀
γc
c
tid'
,
finish_handle
γc
c
Ψ
-∗
WP
f
[
#
c
]
@
tid'
;
⊤
{{
_,
True
}}
}}}
spawn
[
e
]
@
tid
;
⊤
spawn
[
e
]
@
tid
;
⊤
{{{
γc
γe
c
,
RET
#
c
;
join_handle
γc
γe
c
Ψ
}}}
.
{{{
γc
c
,
RET
#
c
;
join_handle
γc
c
Ψ
}}}
.
Proof
.
Proof
.
iIntros
(
<-
Φ
)
"Hf HΦ"
.
rewrite
/
spawn
/=.
iIntros
(
<-
Φ
)
"Hf HΦ"
.
rewrite
/
spawn
/=.
wp_let
.
wp_alloc
l
as
"Hm"
"Hl"
"H†"
.
wp_let
.
wp_let
.
wp_alloc
c
as
"Hm"
"Hl"
"H†"
.
wp_let
.
iMod
(
own_alloc
(
Excl
()))
as
(
γe
)
"Hγe"
;
first
done
.
rewrite
own_loc_na_vec_cons
own_loc_na_vec_singleton
.
rewrite
own_loc_na_vec_cons
own_loc_na_vec_singleton
.
iDestruct
"Hl"
as
"[Hc Hd]"
.
wp_write
.
iDestruct
"Hl"
as
"[Hc Hd]"
.
wp_write
.
iMod
(
GPS_vSP_Init
N
(
λ
γi
,
spawn_interp
γe
γi
Ψ
)
(
λ
_,
IW
)
_
()
with
"Hc []"
)
iMod
(
AtomicPtsTo_from_na
with
"Hc"
)
as
(
γc
t
Vc
)
"(sVc & SW & SP)"
.
as
(
γi
γc
t
)
"[[Htok1 Htok2] SW]"
.
{
iIntros
(???)
.
by
iExists
false
.
}
iMod
(
view_inv_alloc
N
)
as
(
γi
)
"VS"
.
iDestruct
(
GPS_vSP_SWWriter_Reader
with
"SW"
)
as
"#R"
.
iMod
(
"VS"
$!
(
mp_inv_reclaim'
γc
γi
c
Ψ
)
with
"[SP]"
)
as
"[#Inv [Htok1 Htok2]]"
.
{
iNext
.
rewrite
mp_inv_reclaim'_eq
.
iExists
_,
false
,
t
,
Vc
.
by
iFrame
.
}
iDestruct
(
AtomicSWriter_AtomicSeen
with
"SW"
)
as
"#R"
.
wp_apply
(
wp_fork
with
"[SW Htok1 Hf Hd]"
);
[
done
|..]
.
wp_apply
(
wp_fork
with
"[SW Htok1 Hf Hd]"
);
[
done
|..]
.
-
iNext
.
iIntros
(
tid'
)
.
iApply
"Hf"
.
iExists
_,
_,
_
.
by
iFrame
.
-
iNext
.
iIntros
(
tid'
)
.
iApply
"Hf"
.
iExists
_,
_,
_,
_
.
by
iFrame
.
-
iIntros
"_"
.
wp_seq
.
iApply
(
"HΦ"
$!
γc
γe
)
.
iExists
_,
_,
_
.
-
iIntros
"_"
.
wp_seq
.
iApply
(
"HΦ"
$!
γc
)
.
iExists
_,
_,
_,
_
.
iFrame
"R ∗"
.
auto
.
iFrame
"
Inv
R ∗"
.
auto
.
Qed
.
Qed
.
Lemma
finish_spec
Ψ
c
v
γc
γe
tid
Lemma
finish_spec
Ψ
c
v
γc
tid
(
DISJ1
:
(
↑
histN
)
##
(
↑
N
:
coPset
))
(
DISJ2
:
(
↑
escrowN
)
##
(
↑
N
:
coPset
))
:
(
DISJ1
:
(
↑
histN
)
##
(
↑
N
:
coPset
))
:
{{{
finish_handle
γc
γe
c
Ψ
∗
Ψ
v
}}}
{{{
finish_handle
γc
c
Ψ
∗
Ψ
v
}}}
finish
[
#
c
;
v
]
@
tid
;
⊤
finish
[
#
c
;
v
]
@
tid
;
⊤
{{{
RET
#☠
;
True
}}}
.
{{{
RET
#☠
;
True
}}}
.
Proof
.
Proof
.
iIntros
(
Φ
)
"[Hfin HΨ] HΦ"
.
iIntros
(
Φ
)
"[Hfin HΨ] HΦ"
.
iDestruct
"Hfin"
as
(
γi
t
v0
)
"(Hd & SW & Hf)"
.
iDestruct
"Hfin"
as
(
γi
t
v0
V
)
"(Hd & SW & #Inv & vTok1)"
.
wp_lam
.
wp_op
.
wp_write
.
wp_bind
(_
<-
ʳᵉˡ
_)
%
E
.
wp_lam
.
wp_op
.
wp_write
.
iApply
(
GPS_vSP_SWWrite_rel
N
(
spawn_interp
γe
γi
Ψ
)
IW
_
(
1
/
2
)
%
Qp
iMod
(
view_inv_acc_base'
with
"[$Inv $vTok1]"
)
as
"(vTok1 & INV) {Inv}"
;
[
done
|]
.
(
λ
_,
True
)
%
I
True
%
I
(
spawn_interp
γe
γi
Ψ
true
c
γc
t
tt
#
0
)
iDestruct
"INV"
as
(
Vb
)
"[INV Close]"
.
t
()
()
_
#
1
with
"[$Hf $SW HΨ Hd]"
);
[
done
..|
|]
.
rewrite
{
1
}
mp_inv_reclaim'_eq
.
{
iSplitR
""
.
iDestruct
"INV"
as
(
ζ'
b
t0
V0
)
"[Pts _]"
.
-
iIntros
"!>"
(
t'
Ext'
)
"SW _ tok !>"
.
iSplitL
""
;
[
by
iIntros
"!> _"
|]
.
rewrite
view_join_later
.
iDestruct
"Pts"
as
">Pts"
.
iSplitR
""
;
[|
done
]
.
iExists
true
.
iDestruct
(
monPred_in_bot
)
as
"#sV0"
.
iSplitL
""
;
[
done
|]
.
iApply
escrow_alloc
;
[
solve_ndisj
|]
.
iFrame
.
iDestruct
(
view_join_elim'
with
"Pts sV0"
)
as
(
V'
)
"(#sV' & _ & Pts)"
.
iSplitL
"SW"
;
iExists
_;
by
iFrame
.
iDestruct
(
AtomicPtsTo_AtomicSWriter_agree_1
with
"Pts SW"
)
as
%->
.
-
by
iIntros
"!> !> $"
.
}
iIntros
"!>"
(
t'
)
"_"
.
wp_seq
.
by
iApply
"HΦ"
.
iApply
(
AtomicSWriter_release_write
_
_
_
_
V'
_
#
1
(
view_tok
γi
(
1
/
2
)
∗
(
c
>>
1
)
↦
v
∗
Ψ
v
)
%
I
with
"[$SW $Pts $Hd $HΨ $vTok1 $sV']"
);
[
solve_ndisj
|..]
.
iIntros
"!>"
(
t1
V1
)
"((%MAX & %LeV1 & _) & SeenV1 & [[vTok1 Hm] SW'] & Pts')"
.
(* reestablishing the invariant *)
rewrite
bi
.
and_elim_r
bi
.
and_elim_l
.
iMod
(
"Close"
$!
V1
True
%
I
with
"vTok1 [-HΦ]"
)
.
{
iIntros
"vTok1 !>"
.
iSplit
;
[|
done
]
.
rewrite
view_at_view_join
.
iNext
.
rewrite
mp_inv_reclaim'_eq
.
iExists
_,
true
,
t
,
_
.
iSplitL
"Pts'"
;
first
by
iFrame
.
iExists
t1
,
V1
.
iSplit
.
{
iPureIntro
.
split
;
[|
done
]
.
apply
MAX
.
rewrite
lookup_insert
.
by
eexists
.
}
iExists
v
.
by
iFrame
.
}
iIntros
"!>"
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
Lemma
join_spec
Ψ
c
γc
γe
tid
Lemma
join_spec
Ψ
c
γc
tid
(
DISJ1
:
(
↑
histN
)
##
(
↑
N
:
coPset
))
(
DISJ2
:
(
↑
escrowN
)
##
(
↑
N
:
coPset
))
:
(
DISJ1
:
(
↑
histN
)
##
(
↑
N
:
coPset
))
:
{{{
join_handle
γc
γe
c
Ψ
}}}
join
[
#
c
]
@
tid
;
⊤
{{{
v
,
RET
v
;
Ψ
v
}}}
.
{{{
join_handle
γc
c
Ψ
}}}
join
[
#
c
]
@
tid
;
⊤
{{{
v
,
RET
v
;
Ψ
v
}}}
.
Proof
.
Proof
.
iIntros
(
Φ
)
"H HΦ"
.
iIntros
(
Φ
)
"H HΦ"
.
iDestruct
"H"
as
(
γi
Ψ'
t
)
"(R & (H† & He) & Hj & #HΨ')"
.
iDestruct
"H"
as
(
γi
Ψ'
t
V
)
"(#R & #Inv & H† & vTok2 & #HΨ')"
.
iLöb
as
"IH"
forall
(
t
)
.
wp_rec
.
wp_bind
(
!
ᵃᶜ
_)
%
E
.
iLöb
as
"IH"
.
wp_rec
.
wp_bind
(
!
ᵃᶜ
_)
%
E
.
set
P
:
vProp
:=
(
⎡
tok
γe
⎤
)
%
I
.
(* open shared invariant *)
iApply
(
GPS_vSP_SWRead_acq_dealloc
N
(
spawn_interp
γe
γi
Ψ'
)
IW
_
P
iMod
(
view_inv_acc_base'
with
"[$Inv $vTok2]"
)
as
"(vTok2 & INV) {Inv}"
;
[
done
|]
.
(
λ
_
_,
▷
∃
v
,
(
c
>>
1
)
↦
{
1
}
v
∗
Ψ'
v
)
%
I
iDestruct
"INV"
as
(
Vb
)
"[INV Close]"
.
(
spawn_interp
γe
γi
Ψ'
false
c
γc
)
(
1
/
2
)
%
Qp
t
()
_
#
true
rewrite
{
1
}
mp_inv_reclaim'_eq
.
with
"[$Hj $R $He]"
);
[
done
..|
|]
.
iDestruct
"INV"
as
(
ζ'
b
t0
V0
)
"[Pts Own]"
.
{
iSplitL
""
.
(* TODO: IntoLater for view join *)
-
iIntros
"!>"
(
t'
[]
v'
[_
Ext'
])
"!>"
.
iSplit
;
last
iSplit
.
rewrite
view_join_later
.
iDestruct
"Pts"
as
">Pts"
.
+
iDestruct
1
as
(
b
)
"[#Eq ES]"
.
iDestruct
(
monPred_in_bot
)
as
"#sV0"
.
destruct
b
;
simpl
;
[
iDestruct
"ES"
as
"#ES"
|];
(* actual read *)
iModIntro
;
iSplitL
""
;
iExists
_;
by
iFrame
"Eq"
.
iApply
(
AtomicSeen_acquire_read_vj
with
"[$Pts $R $sV0]"
);
[
solve_ndisj
|..]
.
+
iDestruct
1
as
(
b
)
"[#Eq ES]"
.
iIntros
"!>"
(
t'
v'
V'
V''
ζ''
)
"(HF & #SV' & SN' & Pts)"
.
destruct
b
;
simpl
;
[
iDestruct
"ES"
as
"#ES"
|];
iDestruct
"HF"
as
%
([
Sub1
Sub2
]
&
Eqt'
&
MAX'
&
LeV''
)
.
iModIntro
;
iSplitL
""
;
iExists
_;
by
iFrame
"Eq"
.
+
iDestruct
1
as
%
?
.
subst
v'
.
iIntros
"!>"
;
iSplit
;
[
done
|
by
iExists
false
]
.
case
(
decide
(
t'
=
t0
))
=>
[?|
NEqt'
]
.
-
iIntros
"!>"
(
t'
[])
"Ext He $"
.
+
subst
t'
.
iDestruct
1
as
(
b
)
"[% Es]"
.
destruct
b
;
last
done
.
(* we must have read the flag to be 0 (i.e. z = 0). *)
iMod
(
escrow_elim
with
"[] Es [$He]"
)
as
"(SW & >$ & $)"
;
[
solve_ndisj
|
|
done
]
.
iAssert
(
⌜
v'
=
#
0
⌝
)
%
I
as
%
Eq0
.
iIntros
"[He1 He2]"
.
iCombine
"He1"
"He2"
as
"He"
.
{
destruct
b
.
iDestruct
(
own_valid
with
"He"
)
as
"%"
.
auto
.
}
-
iDestruct
"Own"
as
(
t1
V1
[
Lt1
Eqζ'
])
"_"
.
iIntros
"!>"
(
t'
[]
v'
)
"(Ext & R & CASE)"
.
iPureIntro
.
case
(
decide
(
v'
=
#
true
))
=>
?
.
rewrite
Eqζ'
in
Sub2
.
apply
(
lookup_weaken
_
_
_
_
Eqt'
)
in
Sub2
.
-
subst
v'
.
iDestruct
"CASE"
as
"[HInv Ho]"
.
iDestruct
"Ho"
as
(
v
)
"[Hd HΨ]"
.
rewrite
lookup_insert_ne
in
Sub2
.
wp_if
.
wp_op
.
wp_read
.
wp_let
.
iApply
(
wp_hist_inv
);
[
done
|]
.
iIntros
"Hist"
.
+
rewrite
lookup_insert
in
Sub2
.
by
inversion
Sub2
.
iMod
(
GPS_INV_dealloc
_
_
_
true
with
"Hist HInv"
)
as
(
t1
s1
v1
)
"[Hc _]"
;
[
done
|]
.
+
clear
-
Lt1
.
intros
?
.
subst
.
lia
.
iAssert
(
c
↦∗
[
v1
;
v
])
%
I
with
"[Hc Hd]"
as
"Hc"
.
-
iDestruct
"Own"
as
%
Eqζ'
.
iPureIntro
.
rewrite
Eqζ'
in
Sub2
.
apply
(
lookup_weaken
_
_
_
_
Eqt'
)
in
Sub2
.
rewrite
lookup_insert
in
Sub2
.
by
inversion
Sub2
.
}
(* then simply keep looping *)
(* first close the invariant *)
rewrite
bi
.
and_elim_l
.
iMod
(
"Close"
with
"vTok2 [Pts Own]"
)
as
"vTok2"
.
{
iClear
"IH"
.
iNext
.
rewrite
mp_inv_reclaim'_eq
.
iExists
ζ'
,
b
,
t0
,
V0
.
iSplitL
"Pts"
;
by
iFrame
.
}
iIntros
"!>"
.
rewrite
Eq0
.
wp_if
.
(* then apply the Löb IH *)
by
iApply
(
"IH"
with
"H† vTok2 HΦ"
)
.
+
destruct
b
;
last
first
.
{
(* b cannot be false *)
iDestruct
"Own"
as
%
Eqζ'
.
exfalso
.
rewrite
Eqζ'
in
Sub2
.
apply
(
lookup_weaken
_
_
_
_
Eqt'
),
lookup_singleton_Some
in
Sub2
as
[]
.
by
apply
NEqt'
.
}
iClear
"IH"
.
(* we read 1, the destruct Own to get data. *)
iDestruct
"Own"
as
(
t1
V1
[
Lt1
Eqζ'
])
"Own"
.
rewrite
Eqζ'
in
Sub2
.
apply
(
lookup_weaken
_
_
_
_
Eqt'
)
in
Sub2
.
have
?
:
t'
=
t1
.
{
case
(
decide
(
t'
=
t1
))
=>
[
//|
NEqt1
]
.
exfalso
.
by
rewrite
!
lookup_insert_ne
//
in
Sub2
.
}
subst
t'
.
rewrite
lookup_insert
in
Sub2
.
inversion
Sub2
.
subst
v'
V'
.
rewrite
view_join_view_at
.
iDestruct
(
view_at_elim
with
"[SV'] Own"
)
as
(
v
)
"(Hc1 & HΨ & vTok1)"
.
{
iApply
(
monPred_in_mono
with
"SV'"
)
.
simpl
.
solve_lat
.
}
(* cancel the invariant *)
iCombine
"vTok1"
"vTok2"
as
"vTok"
.
rewrite
2
!
bi
.
and_elim_r
.
iDestruct
(
"Close"
with
"vTok"
)
as
"[#LeVb >_]"
.
iDestruct
(
view_join_elim
with
"Pts LeVb"
)
as
"Pts"
.
iIntros
"!>"
.
wp_if
.
wp_op
.
wp_read
.
wp_let
.
iApply
wp_hist_inv
;
[
done
|]
.
iIntros
"HINV"
.
iMod
(
AtomicPtsTo_to_na
with
"HINV Pts"
)
as
(
t'
v'
)
"[Hc _]"
;
[
done
|]
.
iAssert
(
c
↦∗
[
v'
;
v
])
%
I
with
"[Hc Hc1]"
as
"Hc"
.
{
rewrite
own_loc_na_vec_cons
own_loc_na_vec_singleton
.
iFrame
.
}
{
rewrite
own_loc_na_vec_cons
own_loc_na_vec_singleton
.
iFrame
.
}
wp_free
;
first
done
.
iApply
"HΦ"
.
iApply
"HΨ'"
.
done
.
wp_free
;
first
done
.
iApply
"HΦ"
.
by
iApply
"HΨ'"
.
-
iDestruct
"CASE"
as
"(Hj & res & He)"
.
iDestruct
"res"
as
(
b
)
"[% _]"
.
subst
v'
.
destruct
b
;
[
done
|]
.
wp_if
.
iApply
(
"IH"
with
"R H† He Hj"
)
.
auto
.
Qed
.
Qed
.
Lemma
join_handle_impl
Ψ1
Ψ2
γc
γe
c
:
Lemma
join_handle_impl
Ψ1
Ψ2
γc
c
:
□
(
∀
v
,
Ψ1
v
-∗
Ψ2
v
)
-∗
join_handle
γc
γe
c
Ψ1
-∗
join_handle
γc
γe
c
Ψ2
.
□
(
∀
v
,
Ψ1
v
-∗
Ψ2
v
)
-∗
join_handle
γc
c
Ψ1
-∗
join_handle
γc
c
Ψ2
.
Proof
.
Proof
.
iIntros
"#HΨ Hdl"
.
iDestruct
"Hdl"
as
(
γi
Ψ'
t
)
"(? &
(
? & ?
)
& ? & #HΨ')"
.
iIntros
"#HΨ Hdl"
.
iDestruct
"Hdl"
as
(
γi
Ψ'
t
V
)
"(? & ? & ? & ? & #HΨ')"
.
iExists
γi
,
Ψ'
,
t
.
iFrame
"#∗"
.
iIntros
"!# * ?"
.
iExists
γi
,
Ψ'
,
t
,
_
.
iFrame
"#∗"
.
iIntros
"!# * ?"
.
iApply
"HΨ"
.
iApply
"HΨ'"
.
done
.
iApply
"HΨ"
.
iApply
"HΨ'"
.
done
.
Qed
.
Qed
.
End
proof
.
End
proof
.
...
...
This diff is collapsed.
Click to expand it.
theories/typing/lib/join.v
+
2
−
2
View file @
e189f4ad
...
@@ -51,7 +51,7 @@ Section join.
...
@@ -51,7 +51,7 @@ Section join.
iApply
((
spawn_spec
joinN
(
λ
v
,
(
box
R_A
).(
ty_own
)
tid
[
v
]
∗
(
qϝ1
).[
ϝ
])
%
I
)
with
"[HfA HenvA Hϝ1]"
);
iApply
((
spawn_spec
joinN
(
λ
v
,
(
box
R_A
).(
ty_own
)
tid
[
v
]
∗
(
qϝ1
).[
ϝ
])
%
I
)
with
"[HfA HenvA Hϝ1]"
);
first
simpl_subst
.
first
simpl_subst
.
{
(* The new thread. *)
{
(* The new thread. *)
iIntros
(
γc
γe
c
tid'
)
"Hfin"
.
iMod
na_alloc
as
(
tid2
)
"Htl"
.
wp_let
.
wp_let
.
unlock
.
iIntros
(
γc
c
tid'
)
"Hfin"
.
iMod
na_alloc
as
(
tid2
)
"Htl"
.
wp_let
.
wp_let
.
unlock
.
rewrite
!
tctx_hasty_val
.
rewrite
!
tctx_hasty_val
.
iApply
(
type_call_iris
_
[
ϝ
]
()
[_]
_
(
Build_thread_id
tid2
tid'
)
with
"LFT HE Htl [Hϝ1] HfA [HenvA]"
)
.
iApply
(
type_call_iris
_
[
ϝ
]
()
[_]
_
(
Build_thread_id
tid2
tid'
)
with
"LFT HE Htl [Hϝ1] HfA [HenvA]"
)
.
-
rewrite
outlives_product
.
solve_typing
.
-
rewrite
outlives_product
.
solve_typing
.
...
@@ -60,7 +60,7 @@ Section join.
...
@@ -60,7 +60,7 @@ Section join.
-
iIntros
(
r
)
"Htl Hϝ1 Hret"
.
-
iIntros
(
r
)
"Htl Hϝ1 Hret"
.
wp_rec
.
iApply
(
finish_spec
with
"[$Hfin Hret Hϝ1]"
);
[
solve_ndisj
..|
|
auto
]
.
wp_rec
.
iApply
(
finish_spec
with
"[$Hfin Hret Hϝ1]"
);
[
solve_ndisj
..|
|
auto
]
.
rewrite
right_id
.
iFrame
.
by
iApply
@
send_change_tid
.
}
rewrite
right_id
.
iFrame
.
by
iApply
@
send_change_tid
.
}
iNext
.
iIntros
(
γc
γe
c
)
"Hjoin"
.
wp_let
.
wp_let
.
iNext
.
iIntros
(
γc
c
)
"Hjoin"
.
wp_let
.
wp_let
.
iMod
(
lctx_lft_alive_tok_noend
ϝ
with
"HE HL"
)
as
(
qϝ2
)
"(Hϝ2 & HL & Hclose2)"
;
iMod
(
lctx_lft_alive_tok_noend
ϝ
with
"HE HL"
)
as
(
qϝ2
)
"(Hϝ2 & HL & Hclose2)"
;
[
solve_typing
..|]
.
[
solve_typing
..|]
.
rewrite
!
tctx_hasty_val
.
rewrite
!
tctx_hasty_val
.
...
...
This diff is collapsed.
Click to expand it.
theories/typing/lib/spawn.v
+
5
−
5
View file @
e189f4ad
...
@@ -16,7 +16,7 @@ Section join_handle.
...
@@ -16,7 +16,7 @@ Section join_handle.
{|
ty_size
:=
1
;
{|
ty_size
:=
1
;
ty_own
_
vl
:=
ty_own
_
vl
:=
match
vl
return
_
with
match
vl
return
_
with
|
[
#
(
LitLoc
l
)
]
=>
∃
γc
γe
,
join_handle
spawnN
γc
γe
l
(
join_inv
ty
)
|
[
#
(
LitLoc
l
)
]
=>
∃
γc
,
join_handle
spawnN
γc
l
(
join_inv
ty
)
|
_
=>
False
|
_
=>
False
end
%
I
;
end
%
I
;
ty_shr
κ
_
l
:=
True
%
I
|}
.
ty_shr
κ
_
l
:=
True
%
I
|}
.
...
@@ -32,7 +32,7 @@ Section join_handle.
...
@@ -32,7 +32,7 @@ Section join_handle.
Proof
.
Proof
.
iIntros
"#Hincl"
.
iSplit
;
first
done
.
iSplit
;
iModIntro
.
iIntros
"#Hincl"
.
iSplit
;
first
done
.
iSplit
;
iModIntro
.
-
iIntros
"* Hvl"
.
destruct
vl
as
[|[[|
vl
|]|]
[|]];
try
done
.
-
iIntros
"* Hvl"
.
destruct
vl
as
[|[[|
vl
|]|]
[|]];
try
done
.
simpl
.
iDestruct
"Hvl"
as
(
γc
γe
)
"Hvl"
.
iExists
γc
,
γe
.
simpl
.
iDestruct
"Hvl"
as
(
γc
)
"Hvl"
.
iExists
γc
.
iApply
(
join_handle_impl
with
"[] Hvl"
)
.
iIntros
"!> * Hown"
(
tid'
)
.
iApply
(
join_handle_impl
with
"[] Hvl"
)
.
iIntros
"!> * Hown"
(
tid'
)
.
iDestruct
(
box_type_incl
with
"Hincl"
)
as
"{Hincl} (_ & Hincl & _)"
.
iDestruct
(
box_type_incl
with
"Hincl"
)
as
"{Hincl} (_ & Hincl & _)"
.
iApply
"Hincl"
.
done
.
iApply
"Hincl"
.
done
.
...
@@ -94,8 +94,8 @@ Section spawn.
...
@@ -94,8 +94,8 @@ Section spawn.
iApply
(
spawn_spec
_
(
join_inv
retty
)
with
"[-]"
);
iApply
(
spawn_spec
_
(
join_inv
retty
)
with
"[-]"
);
last
first
;
last
simpl_subst
.
last
first
;
last
simpl_subst
.
{
iIntros
"!> *"
.
rewrite
tctx_interp_singleton
tctx_hasty_val
.
{
iIntros
"!> *"
.
rewrite
tctx_interp_singleton
tctx_hasty_val
.
iIntros
"?"
.
iExists
_,
_
.
by
iFrame
.
}
iIntros
"?"
.
iExists
_
.
by
iFrame
.
}
iIntros
(
γc
γe
c
tid'
)
"Hfin"
.
iMod
na_alloc
as
(
tid2
)
"Htl"
.
wp_let
.
wp_let
.
unlock
.
iIntros
(
γc
c
tid'
)
"Hfin"
.
iMod
na_alloc
as
(
tid2
)
"Htl"
.
wp_let
.
wp_let
.
unlock
.
iApply
(
type_call_iris
_
[]
()
[_]
_
(
Build_thread_id
tid2
tid'
)
with
"LFT HE Htl [] Hf' [Henv]"
);
iApply
(
type_call_iris
_
[]
()
[_]
_
(
Build_thread_id
tid2
tid'
)
with
"LFT HE Htl [] Hf' [Henv]"
);
(* The `solve_typing` here shows that, because we assume that `fty` and `retty`
(* The `solve_typing` here shows that, because we assume that `fty` and `retty`
outlive `static`, the implicit requirmeents made by `call_once` are satisifed. *)
outlive `static`, the implicit requirmeents made by `call_once` are satisifed. *)
...
@@ -126,7 +126,7 @@ Section spawn.
...
@@ -126,7 +126,7 @@ Section spawn.
(
λ
r
,
[
r
◁
box
retty
]));
try
solve_typing
;
[|]
.
(
λ
r
,
[
r
◁
box
retty
]));
try
solve_typing
;
[|]
.
{
iIntros
(
tid
qmax
)
"#LFT _ $ $"
.
{
iIntros
(
tid
qmax
)
"#LFT _ $ $"
.
rewrite
tctx_interp_singleton
tctx_hasty_val
.
iIntros
"Hc"
.
rewrite
tctx_interp_singleton
tctx_hasty_val
.
iIntros
"Hc"
.
destruct
c'
as
[[|
c'
|]|];
try
done
.
iDestruct
"Hc"
as
(?
?
)
"Hc"
.
destruct
c'
as
[[|
c'
|]|];
try
done
.
iDestruct
"Hc"
as
(?)
"Hc"
.
iApply
(
join_spec
with
"Hc"
);
[
solve_ndisj
..|]
.
iNext
.
iIntros
"* Hret"
.
iApply
(
join_spec
with
"Hc"
);
[
solve_ndisj
..|]
.
iNext
.
iIntros
"* Hret"
.
rewrite
tctx_interp_singleton
tctx_hasty_val
.
done
.
}
rewrite
tctx_interp_singleton
tctx_hasty_val
.
done
.
}
iIntros
(
r
);
simpl_subst
.
iApply
type_delete
;
[
solve_typing
..|]
.
iIntros
(
r
);
simpl_subst
.
iApply
type_delete
;
[
solve_typing
..|]
.
...
...
This diff is collapsed.
Click to expand it.
Preview
0%
Loading
Try again
or
attach a new file
.
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Save comment
Cancel
Please
register
or
sign in
to comment