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
Simon Spies
Iris
Commits
41bd7cb2
Commit
41bd7cb2
authored
Feb 24, 2016
by
Robbert Krebbers
Browse files
More barrier clean up.
parent
3f8273a4
Changes
2
Expand all
Hide whitespace changes
Inline
Side-by-side
barrier/barrier.v
View file @
41bd7cb2
This diff is collapsed.
Click to expand it.
barrier/client.v
View file @
41bd7cb2
...
...
@@ -5,22 +5,19 @@ Import uPred.
Definition
client
:
=
(
let
:
"b"
:
=
newchan
'
()
in
wait
"b"
)%
L
.
Section
client
.
Context
{
Σ
:
iFunctorG
}
`
{!
heapG
Σ
}
`
{!
barrierG
Σ
}.
Context
(
N
:
namespace
)
(
heapN
:
namespace
).
Context
{
Σ
:
iFunctorG
}
`
{!
heapG
Σ
,
!
barrierG
Σ
}
(
heapN
N
:
namespace
).
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Lemma
client_safe
:
heapN
⊥
N
→
heap_ctx
heapN
⊑
||
client
{{
λ
_
,
True
}}.
Proof
.
intros
?.
rewrite
/
client
.
ewp
eapply
(
newchan_spec
N
heapN
True
%
I
)
;
last
done
.
ewp
eapply
(
newchan_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
.
apply
sep_mono_r
.
apply
wand_intro_r
.
eauto
.
apply
sep_mono_r
,
wand_intro_r
.
eauto
.
Qed
.
End
client
.
Section
ClosedProofs
.
...
...
@@ -33,7 +30,7 @@ Section ClosedProofs.
{
(* FIXME Really?? set_solver takes forever on "⊆ ⊤"?!? *)
by
move
=>?
_
.
}
apply
wp_strip_pvs
,
exist_elim
=>
?.
rewrite
and_elim_l
.
rewrite
-(
client_safe
(
nroot
.@
"
Heap"
)
(
nroot
.@
"
Barrier"
))
//.
rewrite
-(
client_safe
(
nroot
.@
"
Barrier"
)
(
nroot
.@
"
Heap"
))
//.
(* This, too, should be automated. *)
by
apply
ndot_ne_disjoint
.
Qed
.
...
...
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