Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Iris
examples
Commits
6cfb76f6
Commit
6cfb76f6
authored
Feb 22, 2018
by
Jacques-Henri Jourdan
Browse files
Bump iris, fix build.
parent
7fbf8e44
Changes
3
Show whitespace changes
Inline
Side-by-side
opam
View file @
6cfb76f6
...
...
@@ -9,5 +9,5 @@ build: [make "-j%{jobs}%"]
install: [make "install"]
remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris_examples"]
depends: [
"coq-iris" { (= "dev.201
7-10
-2
8
.1
0
") | (= "dev") }
"coq-iris" { (= "dev.201
8-02
-2
1
.1") | (= "dev") }
]
theories/barrier/example_client.v
View file @
6cfb76f6
...
...
@@ -61,7 +61,7 @@ Section ClosedProofs.
Let
Σ
:
gFunctors
:
=
#[
heap
Σ
;
barrier
Σ
;
spawn
Σ
].
Lemma
client_adequate
σ
:
adequate
client
σ
(
λ
_
,
True
).
Lemma
client_adequate
σ
:
adequate
NotStuck
client
σ
(
λ
_
,
True
).
Proof
.
apply
(
heap_adequacy
Σ
)=>
?.
apply
client_safe
.
Qed
.
End
ClosedProofs
.
...
...
theories/barrier/proof.v
View file @
6cfb76f6
...
...
@@ -10,9 +10,9 @@ Set Default Proof Using "Type".
(** The CMRAs/functors we need. *)
Class
barrierG
Σ
:
=
BarrierG
{
barrier_stsG
:
>
stsG
Σ
sts
;
barrier_savedPropG
:
>
savedPropG
Σ
idCF
;
barrier_savedPropG
:
>
savedPropG
Σ
;
}.
Definition
barrier
Σ
:
gFunctors
:
=
#[
sts
Σ
sts
;
savedProp
Σ
idCF
].
Definition
barrier
Σ
:
gFunctors
:
=
#[
sts
Σ
sts
;
savedProp
Σ
].
Instance
subG_barrier
Σ
{
Σ
}
:
subG
barrier
Σ
Σ
→
barrierG
Σ
.
Proof
.
solve_inG
.
Qed
.
...
...
@@ -94,7 +94,7 @@ Proof.
iIntros
(
Φ
)
"_ HΦ"
.
rewrite
-
wp_fupd
/
newbarrier
/=.
wp_seq
.
wp_alloc
l
as
"Hl"
.
iApply
(
"HΦ"
with
"[> -]"
).
iMod
(
saved_prop_alloc
(
F
:
=
idCF
)
P
)
as
(
γ
)
"#?"
.
iMod
(
saved_prop_alloc
P
)
as
(
γ
)
"#?"
.
iMod
(
sts_alloc
(
barrier_inv
l
P
)
_
N
(
State
Low
{[
γ
]})
with
"[-]"
)
as
(
γ
'
)
"[#? Hγ']"
;
eauto
.
{
iNext
.
rewrite
/
barrier_inv
/=.
iFrame
.
...
...
@@ -165,8 +165,8 @@ Proof.
iIntros
(?).
iDestruct
1
as
(
γ
P
Q
i
)
"(#Hsts & Hγ & #HQ & HQR)"
.
iMod
(
sts_openS
(
barrier_inv
l
P
)
_
_
γ
with
"[Hγ]"
)
as
([
p
I
])
"(% & [Hl Hr] & Hclose)"
;
eauto
.
iMod
(
saved_prop_alloc_strong
(
R1
:
∙
%
CF
(
iProp
Σ
))
I
)
as
(
i1
)
"[% #Hi1]"
.
iMod
(
saved_prop_alloc_strong
(
R2
:
∙
%
CF
(
iProp
Σ
))
(
I
∪
{[
i1
]}))
iMod
(
saved_prop_alloc_strong
I
)
as
(
i1
)
"[% #Hi1]"
.
iMod
(
saved_prop_alloc_strong
(
I
∪
{[
i1
]}))
as
(
i2
)
"[Hi2' #Hi2]"
;
iDestruct
"Hi2'"
as
%
Hi2
.
rewrite
->
not_elem_of_union
,
elem_of_singleton
in
Hi2
;
destruct
Hi2
.
iMod
(
"Hclose"
$!
(
State
p
({[
i1
;
i2
]}
∪
I
∖
{[
i
]}))
...
...
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment