Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Iris
Iris
Commits
90734cf0
Commit
90734cf0
authored
Feb 28, 2016
by
Ralf Jung
Browse files
newchan -> newbarrier
parent
17736977
Pipeline
#188
passed with stage
Changes
4
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
barrier/barrier.v
View file @
90734cf0
From
heap_lang
Require
Export
substitution
notation
.
Definition
new
chan
:
val
:
=
λ
:
""
,
ref
'
0
.
Definition
new
barrier
:
val
:
=
λ
:
""
,
ref
'
0
.
Definition
signal
:
val
:
=
λ
:
"x"
,
"x"
<-
'
1
.
Definition
wait
:
val
:
=
rec
:
"wait"
"x"
:
=
if
:
!
"x"
=
'
1
then
'
()
else
"wait"
"x"
.
Instance
new
chan
_closed
:
Closed
new
chan
.
Proof
.
solve_closed
.
Qed
.
Instance
new
barrier
_closed
:
Closed
new
barrier
.
Proof
.
solve_closed
.
Qed
.
Instance
signal_closed
:
Closed
signal
.
Proof
.
solve_closed
.
Qed
.
Instance
wait_closed
:
Closed
wait
.
Proof
.
solve_closed
.
Qed
.
\ No newline at end of file
Instance
wait_closed
:
Closed
wait
.
Proof
.
solve_closed
.
Qed
.
barrier/client.v
View file @
90734cf0
...
...
@@ -2,7 +2,7 @@ From barrier Require Import proof.
From
program_logic
Require
Import
auth
sts
saved_prop
hoare
ownership
.
Import
uPred
.
Definition
client
:
=
(
let
:
"b"
:
=
new
chan
'
()
in
wait
"b"
)%
L
.
Definition
client
:
=
(
let
:
"b"
:
=
new
barrier
'
()
in
wait
"b"
)%
L
.
Section
client
.
Context
{
Σ
:
iFunctorG
}
`
{!
heapG
Σ
,
!
barrierG
Σ
}
(
heapN
N
:
namespace
).
...
...
@@ -12,7 +12,7 @@ Section client.
heapN
⊥
N
→
heap_ctx
heapN
⊑
||
client
{{
λ
_
,
True
}}.
Proof
.
intros
?.
rewrite
/
client
.
ewp
eapply
(
new
chan
_spec
heapN
N
True
%
I
)
;
last
done
.
ewp
eapply
(
new
barrier
_spec
heapN
N
True
%
I
)
;
last
done
.
apply
sep_intro_True_r
;
first
done
.
apply
forall_intro
=>
l
.
apply
wand_intro_l
.
rewrite
right_id
.
wp_let
.
etrans
;
last
eapply
wait_spec
.
...
...
barrier/proof.v
View file @
90734cf0
...
...
@@ -118,12 +118,12 @@ Proof.
Qed
.
(** Actual proofs *)
Lemma
new
chan
_spec
(
P
:
iProp
)
(
Φ
:
val
→
iProp
)
:
Lemma
new
barrier
_spec
(
P
:
iProp
)
(
Φ
:
val
→
iProp
)
:
heapN
⊥
N
→
(
heap_ctx
heapN
★
∀
l
,
recv
l
P
★
send
l
P
-
★
Φ
(
LocV
l
))
⊑
||
new
chan
'
()
{{
Φ
}}.
⊑
||
new
barrier
'
()
{{
Φ
}}.
Proof
.
intros
HN
.
rewrite
/
new
chan
.
wp_seq
.
intros
HN
.
rewrite
/
new
barrier
.
wp_seq
.
rewrite
-
wp_pvs
.
wp
eapply
wp_alloc
;
eauto
with
I
ndisj
.
apply
forall_intro
=>
l
.
rewrite
(
forall_elim
l
).
apply
wand_intro_l
.
rewrite
!
assoc
.
apply
pvs_wand_r
.
...
...
barrier/specification.v
View file @
90734cf0
...
...
@@ -12,7 +12,7 @@ Local Notation iProp := (iPropG heap_lang Σ).
Lemma
barrier_spec
(
heapN
N
:
namespace
)
:
heapN
⊥
N
→
∃
recv
send
:
loc
→
iProp
-
n
>
iProp
,
(
∀
P
,
heap_ctx
heapN
⊑
{{
True
}}
new
chan
'
()
{{
λ
v
,
∃
l
,
v
=
LocV
l
★
recv
l
P
★
send
l
P
}})
∧
(
∀
P
,
heap_ctx
heapN
⊑
{{
True
}}
new
barrier
'
()
{{
λ
v
,
∃
l
,
v
=
LocV
l
★
recv
l
P
★
send
l
P
}})
∧
(
∀
l
P
,
{{
send
l
P
★
P
}}
signal
(
LocV
l
)
{{
λ
_
,
True
}})
∧
(
∀
l
P
,
{{
recv
l
P
}}
wait
(
LocV
l
)
{{
λ
_
,
P
}})
∧
(
∀
l
P
Q
,
{{
recv
l
(
P
★
Q
)
}}
Skip
{{
λ
_
,
recv
l
P
★
recv
l
Q
}})
∧
...
...
@@ -22,7 +22,7 @@ Proof.
exists
(
λ
l
,
CofeMor
(
recv
heapN
N
l
)),
(
λ
l
,
CofeMor
(
send
heapN
N
l
)).
split_and
?
;
simpl
.
-
intros
P
.
apply
:
always_intro
.
apply
impl_intro_r
.
rewrite
-(
new
chan
_spec
heapN
N
P
)
//
always_and_sep_r
.
rewrite
-(
new
barrier
_spec
heapN
N
P
)
//
always_and_sep_r
.
apply
sep_mono_r
,
forall_intro
=>
l
;
apply
wand_intro_l
.
by
rewrite
right_id
-(
exist_intro
l
)
const_equiv
//
left_id
.
-
intros
l
P
.
apply
ht_alt
.
by
rewrite
-
signal_spec
right_id
.
...
...
Write
Preview
Supports
Markdown
0%
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!
Cancel
Please
register
or
sign in
to comment