Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
E
examples
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
0
Issues
0
List
Boards
Labels
Service Desk
Milestones
Merge Requests
1
Merge Requests
1
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Snippets
Snippets
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
Open sidebar
Iris
examples
Commits
65dcfa79
Commit
65dcfa79
authored
Dec 05, 2018
by
Daniel Gratzer
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
New concurrent stack formalization
parent
24efd927
Changes
5
Hide whitespace changes
Inline
Side-by-side
Showing
5 changed files
with
974 additions
and
1279 deletions
+974
-1279
theories/concurrent_stacks/concurrent_stack1.v
theories/concurrent_stacks/concurrent_stack1.v
+137
-149
theories/concurrent_stacks/concurrent_stack2.v
theories/concurrent_stacks/concurrent_stack2.v
+302
-325
theories/concurrent_stacks/concurrent_stack3.v
theories/concurrent_stacks/concurrent_stack3.v
+154
-188
theories/concurrent_stacks/concurrent_stack4.v
theories/concurrent_stacks/concurrent_stack4.v
+381
-583
theories/concurrent_stacks/spec.v
theories/concurrent_stacks/spec.v
+0
-34
No files found.
theories/concurrent_stacks/concurrent_stack1.v
View file @
65dcfa79
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
base_logic
.
lib
Require
Import
invariants
.
From
iris
.
base_logic
.
lib
Require
Import
invariants
.
From
iris
.
program_logic
Require
Export
weakestpre
hoare
.
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris
.
heap_lang
Require
Export
lang
.
From
iris
.
heap_lang
Require
Import
notation
proofmode
.
From
iris
.
algebra
Require
Import
agree
list
.
From
iris
.
heap_lang
Require
Import
assert
proofmode
notation
.
From
iris_examples
.
concurrent_stacks
Require
Import
spec
.
Set
Default
Proof
Using
"Type"
.
Set
Default
Proof
Using
"Type"
.
(** Stack 1: No helping, bag spec. *)
Definition
new_stack
:
val
:
=
λ
:
"_"
,
ref
NONEV
.
Definition
push
:
val
:
=
Definition
mk_stack
:
val
:
=
rec
:
"push"
"s"
"v"
:
=
λ
:
"_"
,
let
:
"tail"
:
=
!
"s"
in
let
:
"r"
:
=
ref
NONEV
in
let
:
"new"
:
=
SOME
(
ref
(
"v"
,
"tail"
))
in
(
rec
:
"pop"
"n"
:
=
if
:
CAS
"s"
"tail"
"new"
then
#()
else
"push"
"s"
"v"
.
match
:
!
"r"
with
Definition
pop
:
val
:
=
NONE
=>
NONE
rec
:
"pop"
"s"
:
=
|
SOME
"hd"
=>
match
:
!
"s"
with
if
:
CAS
"r"
(
SOME
"hd"
)
(
Snd
!
"hd"
)
NONE
=>
NONEV
then
SOME
(
Fst
!
"hd"
)
|
SOME
"l"
=>
else
"pop"
"n"
let
:
"pair"
:
=
!
"l"
in
end
,
if
:
CAS
"s"
(
SOME
"l"
)
(
Snd
"pair"
)
rec
:
"push"
"n"
:
=
then
SOME
(
Fst
"pair"
)
let
:
"r'"
:
=
!
"r"
in
else
"pop"
"s"
let
:
"r''"
:
=
SOME
(
Alloc
(
"n"
,
"r'"
))
in
end
.
if
:
CAS
"r"
"r'"
"r''"
then
#()
else
"push"
"n"
).
Section
stacks
.
Section
stacks
.
Context
`
{!
heapG
Σ
}.
Context
`
{!
heapG
Σ
}
(
N
:
namespace
)
.
Implicit
Types
l
:
loc
.
Implicit
Types
l
:
loc
.
Definition
oloc_to_val
(
h
:
option
loc
)
:
val
:
=
Local
Notation
"l ↦{-} v"
:
=
(
∃
q
,
l
↦
{
q
}
v
)%
I
match
h
with
(
at
level
20
,
format
"l ↦{-} v"
)
:
bi_scope
.
|
None
=>
NONEV
|
Some
l
=>
SOMEV
#
l
Lemma
partial_mapsto_duplicable
l
v
:
end
.
l
↦
{-}
v
-
∗
l
↦
{-}
v
∗
l
↦
{-}
v
.
Local
Instance
oloc_to_val_inj
:
Inj
(=)
(=)
oloc_to_val
.
Proof
.
Proof
.
rewrite
/
Inj
/
oloc_to_val
=>??.
repeat
case_match
;
congruence
.
Qed
.
iIntros
"H"
;
iDestruct
"H"
as
(?)
"[Hl Hl']"
;
iSplitL
"Hl"
;
eauto
.
Qed
.
Definition
is_stack_pre
(
P
:
val
→
iProp
Σ
)
(
F
:
option
loc
-
c
>
iProp
Σ
)
:
Definition
is_list_pre
(
P
:
val
→
iProp
Σ
)
(
F
:
val
-
c
>
iProp
Σ
)
:
option
loc
-
c
>
iProp
Σ
:
=
λ
v
,
val
-
c
>
iProp
Σ
:
=
λ
v
,
(
match
v
with
(
v
≡
NONEV
∨
∃
(
l
:
loc
)
(
h
t
:
val
),
⌜
v
≡
SOMEV
#
l
⌝
∗
l
↦
{-}
(
h
,
t
)%
V
∗
P
h
∗
▷
F
t
)%
I
.
|
None
=>
True
|
Some
l
=>
∃
q
h
t
,
l
↦
{
q
}
(
h
,
oloc_to_val
t
)
∗
P
h
∗
▷
F
t
end
)%
I
.
Local
Instance
is_
stack_contr
(
P
:
val
→
iProp
Σ
)
:
Contractive
(
is_stack
_pre
P
).
Local
Instance
is_
list_contr
(
P
:
val
→
iProp
Σ
)
:
Contractive
(
is_list
_pre
P
).
Proof
.
Proof
.
rewrite
/
is_
stack
_pre
=>
n
f
f'
Hf
v
.
rewrite
/
is_
list
_pre
=>
n
f
f'
Hf
v
.
repeat
(
f_contractive
||
f_equiv
).
repeat
(
f_contractive
||
f_equiv
).
apply
Hf
.
apply
Hf
.
Qed
.
Qed
.
Definition
is_stack_def
(
P
:
val
->
iProp
Σ
)
:
=
fixpoint
(
is_stack_pre
P
).
Definition
is_list_def
(
P
:
val
->
iProp
Σ
)
:
=
fixpoint
(
is_list_pre
P
).
Definition
is_stack_aux
P
:
seal
(@
is_stack_def
P
).
by
eexists
.
Qed
.
Definition
is_list_aux
P
:
seal
(@
is_list_def
P
).
by
eexists
.
Qed
.
Definition
is_stack
P
:
=
unseal
(
is_stack_aux
P
).
Definition
is_list
P
:
=
unseal
(
is_list_aux
P
).
Definition
is_stack_eq
P
:
@
is_stack
P
=
@
is_stack_def
P
:
=
seal_eq
(
is_stack_aux
P
).
Definition
is_list_eq
P
:
@
is_list
P
=
@
is_list_def
P
:
=
seal_eq
(
is_list_aux
P
).
Definition
stack_inv
P
l
:
=
(
∃
ol
,
l
↦
oloc_to_val
ol
∗
is_stack
P
ol
)%
I
.
Lemma
is_list_unfold
(
P
:
val
→
iProp
Σ
)
v
:
is_list
P
v
⊣
⊢
is_list_pre
P
(
is_list
P
)
v
.
Proof
.
rewrite
is_list_eq
.
apply
(
fixpoint_unfold
(
is_list_pre
P
)).
Qed
.
Lemma
is_stack_unfold
(
P
:
val
→
iProp
Σ
)
v
:
(* TODO: shouldn't have to explicitly return is_list *)
is_stack
P
v
⊣
⊢
is_stack_pre
P
(
is_stack
P
)
v
.
Lemma
is_list_unboxed
(
P
:
val
→
iProp
Σ
)
v
:
is_list
P
v
-
∗
⌜
val_is_unboxed
v
⌝
∗
is_list
P
v
.
Proof
.
Proof
.
rewrite
is_stack_eq
.
apply
(
fixpoint_unfold
(
is_stack_pre
P
)).
iIntros
"Hstack"
;
iSplit
;
last
done
;
iDestruct
(
is_list_unfold
with
"Hstack"
)
as
"[->|Hstack]"
;
last
iDestruct
"Hstack"
as
(
l
h
t
)
"(-> & _)"
;
done
.
Qed
.
Qed
.
Lemma
is_stack_copy
(
P
:
val
→
iProp
Σ
)
ol
:
Lemma
is_list_disj
(
P
:
val
→
iProp
Σ
)
v
:
is_stack
P
ol
-
∗
is_stack
P
ol
∗
is_list
P
v
-
∗
is_list
P
v
∗
(
⌜
v
≡
NONEV
⌝
∨
∃
(
l
:
loc
)
h
t
,
⌜
v
≡
SOMEV
#
l
%
V
⌝
∗
l
↦
{-}
(
h
,
t
)%
V
).
(
match
ol
with
None
=>
True
|
Some
l
=>
∃
q
h
t
,
l
↦
{
q
}
(
h
,
oloc_to_val
t
)
end
).
Proof
.
Proof
.
iIntros
"Hstack"
.
iIntros
"Hstack"
.
iDestruct
(
is_stack_unfold
with
"Hstack"
)
as
"Hstack"
.
destruct
ol
;
last
first
.
iDestruct
(
is_list_unfold
with
"Hstack"
)
as
"[%|Hstack]"
;
simplify_eq
.
-
iSplitL
;
try
iApply
is_stack_unfold
;
auto
.
-
rewrite
is_list_unfold
;
iSplitR
;
[
iLeft
|]
;
eauto
.
-
iDestruct
"Hstack"
as
(
q
h
t
)
"[[Hl1 Hl2] [HP Hrest]]"
.
-
iDestruct
"Hstack"
as
(
l
h
t
)
"(% & Hl & Hlist)"
.
iSplitR
"Hl2"
;
try
iApply
is_stack_unfold
;
simpl
;
eauto
10
with
iFrame
.
iDestruct
(
partial_mapsto_duplicable
with
"Hl"
)
as
"[Hl1 Hl2]"
;
simplify_eq
.
rewrite
(
is_list_unfold
_
(
InjRV
_
))
;
iSplitR
"Hl2"
;
iRight
;
iExists
_
,
_
,
_;
by
iFrame
.
Qed
.
Qed
.
(* Per-element invariant (i.e., bag spec). *)
Definition
stack_inv
P
v
:
=
Theorem
stack_works
P
Φ
:
(
∃
l
v'
,
⌜
v
=
#
l
⌝
∗
l
↦
v'
∗
is_list
P
v'
)%
I
.
▷
(
∀
(
f
₁
f
₂
:
val
),
(
□
WP
f
₁
#()
{{
v
,
(
∃
v'
,
v
≡
SOMEV
v'
∗
P
v'
)
∨
v
≡
NONEV
}})
Definition
is_stack
(
P
:
val
→
iProp
Σ
)
v
:
=
-
∗
(
∀
(
v
:
val
),
□
(
P
v
-
∗
WP
f
₂
v
{{
v
,
True
}}))
inv
N
(
stack_inv
P
v
).
-
∗
Φ
(
f
₁
,
f
₂
)%
V
)%
I
-
∗
WP
mk_stack
#()
{{
Φ
}}.
Theorem
mk_stack_spec
P
:
WP
new_stack
#()
{{
s
,
is_stack
P
s
}}%
I
.
Proof
.
Proof
.
i
Intros
"HΦ"
.
i
Apply
wp_fupd
.
wp_lam
.
wp_lam
.
wp_alloc
l
as
"Hl"
.
wp_alloc
ℓ
as
"Hl"
.
pose
proof
(
nroot
.@
"N"
)
as
N
.
iMod
(
inv_alloc
N
⊤
(
stack_inv
P
#
ℓ
)
with
"[Hl]"
)
as
"Hinv"
.
rewrite
-
wp_fupd
.
{
iNext
;
iExists
ℓ
,
NONEV
;
iFrame
;
iMod
(
inv_alloc
N
_
(
stack_inv
P
l
)
with
"[Hl]"
)
as
"#Hisstack"
.
by
iSplit
;
last
(
iApply
is_list_unfold
;
iLeft
).
}
{
iExists
None
;
iFrame
;
auto
.
done
.
iApply
is_stack_unfold
.
auto
.
}
Qed
.
wp_pures
.
Theorem
push_spec
P
s
v
:
{{{
is_stack
P
s
∗
P
v
}}}
push
s
v
{{{
RET
#()
;
True
}}}.
Proof
.
iIntros
(
Φ
)
"[#Hstack HP] HΦ"
.
iL
ö
b
as
"IH"
.
wp_lam
.
wp_lam
.
wp_bind
(
Load
_
).
iInv
N
as
(
ℓ
v'
)
"(>% & Hl & Hlist)"
"Hclose"
;
subst
.
wp_load
.
iMod
(
"Hclose"
with
"[Hl Hlist]"
)
as
"_"
.
{
iNext
;
iExists
_
,
_;
by
iFrame
.
}
iModIntro
.
wp_let
.
wp_alloc
ℓ
'
as
"Hl'"
.
wp_let
.
wp_bind
(
CAS
_
_
_
).
iInv
N
as
(
ℓ
''
v''
)
"(>% & >Hl & Hlist)"
"Hclose"
;
simplify_eq
.
destruct
(
decide
(
v'
=
v''
))
as
[
->
|].
-
iDestruct
(
is_list_unboxed
with
"Hlist"
)
as
"[>% Hlist]"
.
wp_cas_suc
.
iMod
(
"Hclose"
with
"[HP Hl Hl' Hlist]"
)
as
"_"
.
{
iNext
;
iExists
_
,
(
InjRV
#
ℓ
'
)
;
iFrame
;
iSplit
;
first
done
;
rewrite
(
is_list_unfold
_
(
InjRV
_
)).
iRight
;
iExists
_
,
_
,
_;
iFrame
;
eauto
.
}
iModIntro
.
wp_if
.
by
iApply
"HΦ"
.
-
iDestruct
(
is_list_unboxed
with
"Hlist"
)
as
"[>% Hlist]"
.
wp_cas_fail
.
iMod
(
"Hclose"
with
"[Hl Hlist]"
)
as
"_"
.
{
iNext
;
iExists
_
,
_;
by
iFrame
.
}
iModIntro
.
wp_if
.
iApply
(
"IH"
with
"HP HΦ"
).
Qed
.
Theorem
pop_works
P
s
:
{{{
is_stack
P
s
}}}
pop
s
{{{
ov
,
RET
ov
;
⌜
ov
=
NONEV
⌝
∨
∃
v
,
⌜
ov
=
SOMEV
v
⌝
∗
P
v
}}}.
Proof
.
iIntros
(
Φ
)
"#Hstack HΦ"
.
iL
ö
b
as
"IH"
.
wp_lam
.
wp_bind
(
Load
_
).
iInv
N
as
(
ℓ
v'
)
"(>% & Hl & Hlist)"
"Hclose"
;
subst
.
wp_load
.
iDestruct
(
is_list_disj
with
"Hlist"
)
as
"[Hlist Hdisj]"
.
iMod
(
"Hclose"
with
"[Hl Hlist]"
)
as
"_"
.
{
iNext
;
iExists
_
,
_;
by
iFrame
.
}
iModIntro
.
iModIntro
.
i
Apply
"HΦ
"
.
i
Destruct
"Hdisj"
as
"[-> | Heq]
"
.
-
iIntros
"!#"
.
-
wp_match
.
i
L
ö
b
as
"IH"
.
i
Apply
"HΦ"
;
by
iLeft
.
wp_rec
.
-
iDestruct
"Heq"
as
(
l
h
t
)
"[-> Hl]"
.
wp_
bind
(!
#
l
)%
E
.
wp_
match
.
wp_bind
(
Load
_
)
.
iInv
N
as
"Hstack"
"Hclose"
.
iInv
N
as
(
ℓ
'
v'
)
"(>% & Hl' & Hlist)"
"Hclose"
.
simplify_eq
.
iDestruct
"H
stack"
as
(
v'
)
"[Hl' Hstack]
"
.
iDestruct
"H
l"
as
(
q
)
"Hl
"
.
wp_load
.
wp_load
.
destruct
v'
as
[
l'
|]
;
simpl
;
last
first
.
iMod
(
"Hclose"
with
"[Hl' Hlist]"
)
as
"_"
.
+
iMod
(
"Hclose"
with
"[Hl' Hstack]"
)
as
"_"
.
{
iNext
;
iExists
_
,
_;
by
iFrame
.
}
{
rewrite
/
stack_inv
.
eauto
with
iFrame
.
}
iModIntro
.
wp_match
.
wp_pures
.
by
iRight
.
+
iDestruct
(
is_stack_copy
with
"Hstack"
)
as
"[Hstack Hmy]"
.
iDestruct
"Hmy"
as
(
q
h
t
)
"Hl"
.
iMod
(
"Hclose"
with
"[Hl' Hstack]"
)
as
"_"
.
{
rewrite
/
stack_inv
.
eauto
with
iFrame
.
}
iModIntro
.
wp_match
.
wp_load
.
wp_pures
.
wp_bind
(
CAS
_
_
_
).
iInv
N
as
"Hstack"
"Hclose"
.
iDestruct
"Hstack"
as
(
v''
)
"[Hl'' Hstack]"
.
destruct
(
decide
(
oloc_to_val
v''
=
oloc_to_val
(
Some
l'
)))
as
[->%
oloc_to_val_inj
|
Hne
].
*
simpl
.
wp_cas_suc
.
iDestruct
(
is_stack_unfold
with
"Hstack"
)
as
"Hstack"
.
iDestruct
"Hstack"
as
(
q'
h'
t'
)
"[Hl' [HP Hstack]]"
.
iDestruct
(
mapsto_agree
with
"Hl Hl'"
)
as
%[=
<-
<-%
oloc_to_val_inj
].
iMod
(
"Hclose"
with
"[Hl'' Hstack]"
).
{
iExists
_
.
auto
with
iFrame
.
}
iModIntro
.
wp_if
.
wp_load
.
wp_pures
.
eauto
.
*
simpl
in
Hne
.
wp_cas_fail
.
iMod
(
"Hclose"
with
"[Hl'' Hstack]"
).
{
iExists
v''
;
iFrame
;
auto
.
}
iModIntro
.
wp_if
.
iApply
"IH"
.
-
iIntros
(
v
)
"!# HP"
.
iL
ö
b
as
"IH"
.
wp_rec
.
wp_bind
(!
_
)%
E
.
iInv
N
as
"Hstack"
"Hclose"
.
iDestruct
"Hstack"
as
(
v'
)
"[Hl' Hstack]"
.
wp_load
.
iMod
(
"Hclose"
with
"[Hl' Hstack]"
).
{
iExists
v'
.
iFrame
.
}
iModIntro
.
iModIntro
.
wp_let
.
wp_let
.
wp_proj
.
wp_bind
(
CAS
_
_
_
).
wp_alloc
r''
as
"Hr''"
.
iInv
N
as
(
ℓ
''
v''
)
"(>% & Hl' & Hlist)"
"Hclose"
.
simplify_eq
.
wp_pures
.
wp_bind
(
CAS
_
_
_
).
destruct
(
decide
(
v''
=
InjRV
#
l
))
as
[->
|].
iInv
N
as
"Hstack"
"Hclose"
.
*
rewrite
is_list_unfold
.
iDestruct
"Hstack"
as
(
v''
)
"[Hl'' Hstack]"
.
iDestruct
"Hlist"
as
"[>% | H]"
;
first
done
.
wp_cas
as
->%
oloc_to_val_inj
|
_
.
iDestruct
"H"
as
(
ℓ
'''
h'
t'
)
"(>% & Hl'' & HP & Hlist)"
;
simplify_eq
.
+
destruct
v'
;
by
right
.
iDestruct
"Hl''"
as
(
q'
)
"Hl''"
.
+
iMod
(
"Hclose"
with
"[Hl'' Hr'' HP Hstack]"
).
wp_cas_suc
.
iExists
(
Some
r''
).
iDestruct
(
mapsto_agree
with
"Hl'' Hl"
)
as
"%"
;
simplify_eq
.
iFrame
;
auto
.
iMod
(
"Hclose"
with
"[Hl' Hlist]"
)
as
"_"
.
iNext
.
{
iNext
;
iExists
ℓ
''
,
_;
by
iFrame
.
}
iApply
is_stack_unfold
.
simpl
.
iExists
_
,
_
,
v'
.
iFrame
.
iModIntro
.
iModIntro
.
wp_if
.
wp_if
.
done
.
wp_proj
.
+
iMod
(
"Hclose"
with
"[Hl'' Hstack]"
).
iApply
(
"HΦ"
with
"[HP]"
)
;
iRight
;
iExists
h
;
by
iFrame
.
iExists
v''
;
iFrame
;
auto
.
*
wp_cas_fail
.
iMod
(
"Hclose"
with
"[Hl' Hlist]"
)
as
"_"
.
{
iNext
;
iExists
ℓ
''
,
_;
by
iFrame
.
}
iModIntro
.
iModIntro
.
wp_if
.
wp_if
.
iApply
"IH"
.
iApply
(
"IH"
with
"HΦ"
).
done
.
Qed
.
Qed
.
End
stacks
.
End
stacks
.
Program
Definition
is_concurrent_bag
`
{!
heapG
Σ
}
:
concurrent_bag
Σ
:
=
{|
spec
.
mk_bag
:
=
mk_stack
|}.
Next
Obligation
.
iIntros
(???
P
Φ
)
"_ HΦ"
.
iApply
stack_works
.
iNext
.
iIntros
(
f
₁
f
₂
)
"Hpop Hpush"
.
iApply
"HΦ"
.
iFrame
.
Qed
.
theories/concurrent_stacks/concurrent_stack2.v
View file @
65dcfa79
(** THIS FILE CURRENTLY DOES NOT COMPILE because it has not been ported to the
From
iris
.
heap_lang
Require
Import
notation
proofmode
.
stricter CAS requirements yet. *)
From
iris
.
program_logic
Require
Export
weakestpre
hoare
.
From
iris
.
heap_lang
Require
Export
lang
proofmode
notation
.
From
iris
.
algebra
Require
Import
excl
.
From
iris
.
algebra
Require
Import
excl
.
From
iris
.
base_logic
.
lib
Require
Import
invariants
.
From
iris_examples
.
concurrent_stacks
Require
Import
spec
.
From
iris
.
program_logic
Require
Export
weakestpre
.
Set
Default
Proof
Using
"Type"
.
Set
Default
Proof
Using
"Type"
.
(** Stack 2: With helping, bag spec. *)
Definition
mk_offer
:
val
:
=
Definition
mk_offer
:
val
:
=
λ
:
"v"
,
(
"v"
,
ref
#
0
).
λ
:
"v"
,
(
"v"
,
ref
#
0
).
Definition
revoke_offer
:
val
:
=
Definition
revoke_offer
:
val
:
=
...
@@ -17,49 +11,49 @@ Definition revoke_offer : val :=
...
@@ -17,49 +11,49 @@ Definition revoke_offer : val :=
Definition
take_offer
:
val
:
=
Definition
take_offer
:
val
:
=
λ
:
"v"
,
if
:
CAS
(
Snd
"v"
)
#
0
#
1
then
SOME
(
Fst
"v"
)
else
NONE
.
λ
:
"v"
,
if
:
CAS
(
Snd
"v"
)
#
0
#
1
then
SOME
(
Fst
"v"
)
else
NONE
.
Definition
m
ailbox
:
val
:
=
Definition
m
k_mailbox
:
val
:
=
λ
:
"_"
,
ref
NONEV
.
λ
:
"_"
,
Definition
put
:
val
:
=
let
:
"r"
:
=
ref
NONEV
in
λ
:
"r"
"v"
,
(
rec
:
"put"
"v"
:
=
let
:
"off"
:
=
mk_offer
"v"
in
let
:
"off"
:
=
mk_offer
"v"
in
"r"
<-
SOME
"off"
;;
"r"
<-
SOME
"off"
;;
revoke_offer
"off"
.
revoke_offer
"off"
,
rec
:
"get"
"n"
:
=
Definition
get
:
val
:
=
let
:
"offopt"
:
=
!
"r"
in
λ
:
"r"
,
match
:
"offopt"
with
let
:
"offopt"
:
=
!
"r"
in
NONE
=>
NONE
match
:
"offopt"
with
|
SOME
"x"
=>
take_offer
"x"
NONE
=>
NONE
end
|
SOME
"x"
=>
take_offer
"x"
)
.
end
.
Definition
mk_stack
:
val
:
=
Definition
mk_stack
:
val
:
=
λ
:
"_"
,
(
mk_mailbox
#(),
ref
NONEV
).
λ
:
"_"
,
Definition
push
:
val
:
=
let
:
"mailbox"
:
=
mailbox
#()
in
rec
:
"push"
"p"
"v"
:
=
let
:
"put"
:
=
Fst
"mailbox
"
in
let
:
"mailbox"
:
=
Fst
"p
"
in
let
:
"get"
:
=
Snd
"mailbox
"
in
let
:
"s"
:
=
Snd
"p
"
in
let
:
"r"
:
=
ref
NONEV
in
match
:
put
"mailbox"
"v"
with
(
rec
:
"pop"
"n"
:
=
NONE
=>
#()
match
:
"get"
#()
with
|
SOME
"v'"
=>
NONE
=>
let
:
"tail"
:
=
!
"s"
in
(
match
:
!
"r"
with
let
:
"new"
:
=
SOME
(
ref
(
"v'"
,
"tail"
))
in
NONE
=>
NONE
if
:
CAS
"s"
"tail"
"new"
then
#()
else
"push"
"p"
"v'"
|
SOME
"hd"
=>
end
.
if
:
CAS
"r"
(
SOME
"hd"
)
(
Snd
"hd"
)
Definition
pop
:
val
:
=
then
SOME
(
Fst
"hd"
)
rec
:
"pop"
"p"
:
=
else
"pop"
"n"
let
:
"mailbox"
:
=
Fst
"p"
in
end
)
let
:
"s"
:
=
Snd
"p"
in
|
SOME
"x"
=>
SOME
"x"
match
:
get
"mailbox"
with
end
,
NONE
=>
rec
:
"push"
"n"
:
=
match
:
!
"s"
with
match
:
"put"
"n"
with
NONE
=>
NONEV
NONE
=>
#()
|
SOME
"l"
=>
|
SOME
"n"
=>
let
:
"pair"
:
=
!
"l"
in
let
:
"r'"
:
=
!
"r"
in
if
:
CAS
"s"
(
SOME
"l"
)
(
Snd
"pair"
)
let
:
"r''"
:
=
SOME
(
"n"
,
"r'"
)
in
then
SOME
(
Fst
"pair"
)
if
:
CAS
"r"
"r'"
"r''
"
else
"pop"
"p
"
then
#()
end
else
"push"
"n
"
|
SOME
"x"
=>
SOME
"x
"
end
)
.
end
.
Definition
channelR
:
=
exclR
unitR
.
Definition
channelR
:
=
exclR
unitR
.
Class
channelG
Σ
:
=
{
channel_inG
:
>
inG
Σ
channelR
}.
Class
channelG
Σ
:
=
{
channel_inG
:
>
inG
Σ
channelR
}.
...
@@ -68,350 +62,333 @@ Instance subG_channelΣ {Σ} : subG channelΣ Σ → channelG Σ.
...
@@ -68,350 +62,333 @@ Instance subG_channelΣ {Σ} : subG channelΣ Σ → channelG Σ.
Proof
.
solve_inG
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
side_channel
.
Section
side_channel
.
Context
`
{!
heapG
Σ
,
!
channelG
Σ
}.
Context
`
{!
heapG
Σ
,
!
channelG
Σ
}
(
N
:
namespace
)
.
Implicit
Types
l
:
loc
.
Implicit
Types
l
:
loc
.
Definition
revoke_tok
γ
:
=
own
γ
(
Excl
()).
Definition
stages
γ
(
P
:
val
→
iProp
Σ
)
l
v
:
=
Definition
stages
γ
(
P
:
val
→
iProp
Σ
)
l
v
:
=
((
l
↦
#
0
∗
P
v
)
((
l
↦
#
0
∗
P
v
)
∨
(
l
↦
#
1
)
∨
(
l
↦
#
1
)
∨
(
l
↦
#
2
∗
own
γ
(
Excl
())
))%
I
.
∨
(
l
↦
#
2
∗
revoke_tok
γ
))%
I
.
Definition
is_offer
γ
(
P
:
val
→
iProp
Σ
)
(
v
:
val
)
:
iProp
Σ
:
=
Definition
is_offer
γ
(
P
:
val
→
iProp
Σ
)
(
v
:
val
)
:
iProp
Σ
:
=
(
∃
v'
l
,
⌜
v
=
(
v'
,
#
l
)%
V
⌝
∗
∃
ι
,
inv
ι
(
stages
γ
P
l
v'
))%
I
.
(
∃
v'
l
,
⌜
v
=
(
v'
,
#
l
)%
V
⌝
∗
inv
N
(
stages
γ
P
l
v'
))%
I
.
Definition
mailbox_inv
(
P
:
val
→
iProp
Σ
)
(
v
:
val
)
:
iProp
Σ
:
=
Lemma
mk_offer_works
P
v
:
(
∃
l
,
⌜
v
=
#
l
⌝
∗
(
l
↦
NONEV
∨
(
∃
v'
γ
,
l
↦
SOMEV
v'
∗
is_offer
γ
P
v'
)))%
I
.
{{{
P
v
}}}
mk_offer
v
{{{
o
γ
,
RET
o
;
is_offer
γ
P
o
∗
revoke_tok
γ
}}}.
Proof
.
iIntros
(
Φ
)
"HP HΦ"
.
rewrite
-
wp_fupd
.
wp_lam
.
wp_alloc
l
as
"Hl"
.
iMod
(
own_alloc
(
Excl
()))
as
(
γ
)
"Hγ"
;
first
done
.
iMod
(
inv_alloc
N
_
(
stages
γ
P
l
v
)
with
"[Hl HP]"
)
as
"#Hinv"
.
{
iNext
;
iLeft
;
iFrame
.
}
iModIntro
;
iApply
"HΦ"
;
iFrame
;
iExists
_
,
_;
auto
.
Qed
.
(* A partial specification for revoke that will be useful later *)
(* A partial specification for revoke that will be useful later *)
Lemma
revoke_works
N
γ
P
l
v
:
Lemma
revoke_works
γ
P
v
:
inv
N
(
stages
γ
P
l
v
)
∗
own
γ
(
Excl
())
-
∗
{{{
is_offer
γ
P
v
∗
revoke_tok
γ
}}}
WP
revoke_offer
(
v
,
#
l
)
revoke_offer
v
{{
v'
,
(
∃
v''
:
val
,
⌜
v'
=
InjRV
v''
⌝
∗
P
v''
)
∨
⌜
v'
=
InjLV
#()
⌝
}}.
{{{
v'
,
RET
v'
;
(
∃
v''
:
val
,
⌜
v'
=
InjRV
v''
⌝
∗
P
v''
)
∨
⌜
v'
=
InjLV
#()
⌝
}
}}.
Proof
.
Proof
.
iIntros
"[#Hinv Hγ]"
.
iIntros
(
Φ
)
"[Hinv Hγ] HΦ"
.
iDestruct
"Hinv"
as
(
v'
l
)
"[-> #Hinv]"
.
wp_let
.
wp_let
.
wp_proj
.
wp_bind
(
CAS
_
_
_
).
wp_proj
.
wp_bind
(
CAS
_
_
_
).
iInv
N
as
"Hstages"
"Hclose"
.
iInv
N
as
"Hstages"
"Hclose"
.
iDestruct
"Hstages"
as
"[H | [H | H]]"
.
iDestruct
"Hstages"
as
"[[Hl HP] | [H | [Hl H]]]"
.
-
iDestruct
"H"
as
"[Hl HP]"
.
-
wp_cas_suc
.
wp_cas_suc
.
iMod
(
"Hclose"
with
"[Hl Hγ]"
)
as
"_"
.
iMod
(
"Hclose"
with
"[Hl Hγ]"
).
{
iRight
;
iRight
;
iFrame
.
}
iRight
;
iRight
;
iFrame
.
iModIntro
.
iModIntro
.
wp_if
.
wp_if
.
wp_proj
.
wp_proj
.
by
iApply
"HΦ"
;
iLeft
;
iExists
_;
iSplit
.
iLeft
.
iExists
v
;
iSplit
;
auto
.
-
wp_cas_fail
.
-
wp_cas_fail
.
iMod
(
"Hclose"
with
"[H]"
).
iMod
(
"Hclose"
with
"[H]"
)
as
"_"
.
iRight
;
iLeft
;
auto
.
{
iRight
;
iLeft
;
auto
.
}
iModIntro
.
iModIntro
.
wp_if
.
wp_if
.
iRight
;
auto
.
by
iApply
"HΦ"
;
iRight
.
-
iDestruct
"H"
as
"[Hl H]"
.
-
wp_cas_fail
.
wp_cas_fail
.
iDestruct
(
own_valid_2
with
"H Hγ"
)
as
%[].
by
iDestruct
(
own_valid_2
with
"H Hγ"
)
as
%?.
Qed
.
Qed
.
(* A partial specification for take that will be useful later *)
(* A partial specification for take that will be useful later *)
Lemma
take_works
γ
N
P
v
l
: