Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
10
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Iris
Iris
Commits
6e487530
Commit
6e487530
authored
Jul 19, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Explicit namespace for heap instead of parametrizing over it.
parent
2d44c0ef
Changes
14
Hide whitespace changes
Inline
Side-by-side
Showing
14 changed files
with
91 additions
and
100 deletions
+91
-100
heap_lang/heap.v
heap_lang/heap.v
+24
-24
heap_lang/lib/barrier/proof.v
heap_lang/lib/barrier/proof.v
+3
-5
heap_lang/lib/barrier/specification.v
heap_lang/lib/barrier/specification.v
+5
-5
heap_lang/lib/counter.v
heap_lang/lib/counter.v
+3
-4
heap_lang/lib/lock.v
heap_lang/lib/lock.v
+4
-5
heap_lang/lib/par.v
heap_lang/lib/par.v
+3
-4
heap_lang/lib/spawn.v
heap_lang/lib/spawn.v
+3
-4
heap_lang/proofmode.v
heap_lang/proofmode.v
+12
-13
tests/barrier_client.v
tests/barrier_client.v
+8
-9
tests/heap_lang.v
tests/heap_lang.v
+9
-9
tests/joining_existentials.v
tests/joining_existentials.v
+7
-7
tests/list_reverse.v
tests/list_reverse.v
+3
-3
tests/one_shot.v
tests/one_shot.v
+3
-4
tests/tree_sum.v
tests/tree_sum.v
+4
-4
No files found.
heap_lang/heap.v
View file @
6e487530
...
...
@@ -8,6 +8,7 @@ Import uPred.
a finmap as their state. Or maybe even beyond "as their state", i.e. arbitrary
predicates over finmaps instead of just ownP. *)
Definition
heapN
:
namespace
:
=
nroot
.@
"heap"
.
Definition
heapUR
:
ucmraT
:
=
gmapUR
loc
(
prodR
fracR
(
dec_agreeR
val
)).
(** The CMRA we need. *)
...
...
@@ -22,7 +23,7 @@ Definition to_heap : state → heapUR := fmap (λ v, (1%Qp, DecAgree v)).
Definition
of_heap
:
heapUR
→
state
:
=
omap
(
maybe
DecAgree
∘
snd
).
Section
definitions
.
Context
`
{
i
:
heapG
Σ
}.
Context
`
{
heapG
Σ
}.
Definition
heap_mapsto_def
(
l
:
loc
)
(
q
:
Qp
)
(
v
:
val
)
:
iPropG
heap_lang
Σ
:
=
auth_own
heap_name
{[
l
:
=
(
q
,
DecAgree
v
)
]}.
...
...
@@ -33,12 +34,12 @@ Section definitions.
Definition
heap_inv
(
h
:
heapUR
)
:
iPropG
heap_lang
Σ
:
=
ownP
(
of_heap
h
).
Definition
heap_ctx
(
N
:
namespace
)
:
iPropG
heap_lang
Σ
:
=
auth_ctx
heap_name
N
heap_inv
.
Definition
heap_ctx
:
iPropG
heap_lang
Σ
:
=
auth_ctx
heap_name
heap
N
heap_inv
.
Global
Instance
heap_inv_proper
:
Proper
((
≡
)
==>
(
⊣
⊢
))
heap_inv
.
Proof
.
solve_proper
.
Qed
.
Global
Instance
heap_ctx_persistent
N
:
PersistentP
(
heap_ctx
N
)
.
Global
Instance
heap_ctx_persistent
:
PersistentP
heap_ctx
.
Proof
.
apply
_
.
Qed
.
End
definitions
.
...
...
@@ -53,7 +54,6 @@ Notation "l ↦ v" := (heap_mapsto l 1 v) (at level 20) : uPred_scope.
Section
heap
.
Context
{
Σ
:
gFunctors
}.
Implicit
Types
N
:
namespace
.
Implicit
Types
P
Q
:
iPropG
heap_lang
Σ
.
Implicit
Types
Φ
:
val
→
iPropG
heap_lang
Σ
.
Implicit
Types
σ
:
state
.
...
...
@@ -103,13 +103,13 @@ Section heap.
Hint
Resolve
heap_store_valid
.
(** Allocation *)
Lemma
heap_alloc
N
E
σ
:
authG
heap_lang
Σ
heapUR
→
nclose
N
⊆
E
→
ownP
σ
={
E
}=>
∃
_
:
heapG
Σ
,
heap_ctx
N
∧
[
★
map
]
l
↦
v
∈
σ
,
l
↦
v
.
Lemma
heap_alloc
E
σ
:
authG
heap_lang
Σ
heapUR
→
nclose
heap
N
⊆
E
→
ownP
σ
={
E
}=>
∃
_
:
heapG
Σ
,
heap_ctx
∧
[
★
map
]
l
↦
v
∈
σ
,
l
↦
v
.
Proof
.
intros
.
rewrite
-{
1
}(
from_to_heap
σ
).
etrans
.
{
rewrite
[
ownP
_
]
later_intro
.
apply
(
auth_alloc
(
ownP
∘
of_heap
)
N
E
)
;
auto
using
to_heap_valid
.
}
apply
(
auth_alloc
(
ownP
∘
of_heap
)
heap
N
E
)
;
auto
using
to_heap_valid
.
}
apply
pvs_mono
,
exist_elim
=>
γ
.
rewrite
-(
exist_intro
(
HeapG
_
_
γ
))
/
heap_ctx
;
apply
and_mono_r
.
rewrite
heap_mapsto_eq
/
heap_mapsto_def
/
heap_name
.
...
...
@@ -149,9 +149,9 @@ Section heap.
(** Weakest precondition *)
(* FIXME: try to reduce usage of wp_pvs. We're losing view shifts here. *)
Lemma
wp_alloc
N
E
e
v
Φ
:
to_val
e
=
Some
v
→
nclose
N
⊆
E
→
heap_ctx
N
★
▷
(
∀
l
,
l
↦
v
={
E
}=
★
Φ
(
LitV
(
LitLoc
l
)))
⊢
WP
Alloc
e
@
E
{{
Φ
}}.
Lemma
wp_alloc
E
e
v
Φ
:
to_val
e
=
Some
v
→
nclose
heap
N
⊆
E
→
heap_ctx
★
▷
(
∀
l
,
l
↦
v
={
E
}=
★
Φ
(
LitV
(
LitLoc
l
)))
⊢
WP
Alloc
e
@
E
{{
Φ
}}.
Proof
.
iIntros
(??)
"[#Hinv HΦ]"
.
rewrite
/
heap_ctx
.
iPvs
(
auth_empty
heap_name
)
as
"Hheap"
.
...
...
@@ -166,9 +166,9 @@ Section heap.
iIntros
"Hheap"
.
iApply
"HΦ"
.
by
rewrite
heap_mapsto_eq
/
heap_mapsto_def
.
Qed
.
Lemma
wp_load
N
E
l
q
v
Φ
:
nclose
N
⊆
E
→
heap_ctx
N
★
▷
l
↦
{
q
}
v
★
▷
(
l
↦
{
q
}
v
={
E
}=
★
Φ
v
)
Lemma
wp_load
E
l
q
v
Φ
:
nclose
heap
N
⊆
E
→
heap_ctx
★
▷
l
↦
{
q
}
v
★
▷
(
l
↦
{
q
}
v
={
E
}=
★
Φ
v
)
⊢
WP
Load
(
Lit
(
LitLoc
l
))
@
E
{{
Φ
}}.
Proof
.
iIntros
(?)
"[#Hh [Hl HΦ]]"
.
...
...
@@ -181,9 +181,9 @@ Section heap.
rewrite
of_heap_singleton_op
//.
by
iFrame
.
Qed
.
Lemma
wp_store
N
E
l
v'
e
v
Φ
:
to_val
e
=
Some
v
→
nclose
N
⊆
E
→
heap_ctx
N
★
▷
l
↦
v'
★
▷
(
l
↦
v
={
E
}=
★
Φ
(
LitV
LitUnit
))
Lemma
wp_store
E
l
v'
e
v
Φ
:
to_val
e
=
Some
v
→
nclose
heap
N
⊆
E
→
heap_ctx
★
▷
l
↦
v'
★
▷
(
l
↦
v
={
E
}=
★
Φ
(
LitV
LitUnit
))
⊢
WP
Store
(
Lit
(
LitLoc
l
))
e
@
E
{{
Φ
}}.
Proof
.
iIntros
(??)
"[#Hh [Hl HΦ]]"
.
...
...
@@ -197,9 +197,9 @@ Section heap.
rewrite
of_heap_singleton_op
//
;
eauto
.
by
iFrame
.
Qed
.
Lemma
wp_cas_fail
N
E
l
q
v'
e1
v1
e2
v2
Φ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
v'
≠
v1
→
nclose
N
⊆
E
→
heap_ctx
N
★
▷
l
↦
{
q
}
v'
★
▷
(
l
↦
{
q
}
v'
={
E
}=
★
Φ
(
LitV
(
LitBool
false
)))
Lemma
wp_cas_fail
E
l
q
v'
e1
v1
e2
v2
Φ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
v'
≠
v1
→
nclose
heap
N
⊆
E
→
heap_ctx
★
▷
l
↦
{
q
}
v'
★
▷
(
l
↦
{
q
}
v'
={
E
}=
★
Φ
(
LitV
(
LitBool
false
)))
⊢
WP
CAS
(
Lit
(
LitLoc
l
))
e1
e2
@
E
{{
Φ
}}.
Proof
.
iIntros
(????)
"[#Hh [Hl HΦ]]"
.
...
...
@@ -212,9 +212,9 @@ Section heap.
rewrite
of_heap_singleton_op
//.
by
iFrame
.
Qed
.
Lemma
wp_cas_suc
N
E
l
e1
v1
e2
v2
Φ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
nclose
N
⊆
E
→
heap_ctx
N
★
▷
l
↦
v1
★
▷
(
l
↦
v2
={
E
}=
★
Φ
(
LitV
(
LitBool
true
)))
Lemma
wp_cas_suc
E
l
e1
v1
e2
v2
Φ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
nclose
heap
N
⊆
E
→
heap_ctx
★
▷
l
↦
v1
★
▷
(
l
↦
v2
={
E
}=
★
Φ
(
LitV
(
LitBool
true
)))
⊢
WP
CAS
(
Lit
(
LitLoc
l
))
e1
e2
@
E
{{
Φ
}}.
Proof
.
iIntros
(???)
"[#Hh [Hl HΦ]]"
.
...
...
heap_lang/lib/barrier/proof.v
View file @
6e487530
...
...
@@ -21,8 +21,7 @@ Proof. destruct H as (?&?&?). split; apply _. Qed.
(** Now we come to the Iris part of the proof. *)
Section
proof
.
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
barrierG
Σ
}.
Context
(
heapN
N
:
namespace
).
Context
`
{!
heapG
Σ
,
!
barrierG
Σ
}
(
N
:
namespace
).
Implicit
Types
I
:
gset
gname
.
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
...
...
@@ -42,7 +41,7 @@ Definition barrier_inv (l : loc) (P : iProp) (s : state) : iProp :=
(
l
↦
s
★
ress
(
state_to_prop
s
P
)
(
state_I
s
))%
I
.
Definition
barrier_ctx
(
γ
:
gname
)
(
l
:
loc
)
(
P
:
iProp
)
:
iProp
:
=
(
■
(
heapN
⊥
N
)
★
heap_ctx
heapN
★
sts_ctx
γ
N
(
barrier_inv
l
P
))%
I
.
(
■
(
heapN
⊥
N
)
★
heap_ctx
★
sts_ctx
γ
N
(
barrier_inv
l
P
))%
I
.
Definition
send
(
l
:
loc
)
(
P
:
iProp
)
:
iProp
:
=
(
∃
γ
,
barrier_ctx
γ
l
P
★
sts_ownS
γ
low_states
{[
Send
]})%
I
.
...
...
@@ -96,8 +95,7 @@ Qed.
(** Actual proofs *)
Lemma
newbarrier_spec
(
P
:
iProp
)
(
Φ
:
val
→
iProp
)
:
heapN
⊥
N
→
heap_ctx
heapN
★
(
∀
l
,
recv
l
P
★
send
l
P
-
★
Φ
#
l
)
⊢
WP
newbarrier
#()
{{
Φ
}}.
heap_ctx
★
(
∀
l
,
recv
l
P
★
send
l
P
-
★
Φ
#
l
)
⊢
WP
newbarrier
#()
{{
Φ
}}.
Proof
.
iIntros
(
HN
)
"[#? HΦ]"
.
rewrite
/
newbarrier
.
wp_seq
.
wp_alloc
l
as
"Hl"
.
...
...
heap_lang/lib/barrier/specification.v
View file @
6e487530
...
...
@@ -9,20 +9,20 @@ Context {Σ : gFunctors} `{!heapG Σ} `{!barrierG Σ}.
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Lemma
barrier_spec
(
heapN
N
:
namespace
)
:
Lemma
barrier_spec
(
N
:
namespace
)
:
heapN
⊥
N
→
∃
recv
send
:
loc
→
iProp
-
n
>
iProp
,
(
∀
P
,
heap_ctx
heapN
⊢
{{
True
}}
newbarrier
#()
{{
v
,
∃
l
:
loc
,
v
=
#
l
★
recv
l
P
★
send
l
P
}})
∧
(
∀
P
,
heap_ctx
⊢
{{
True
}}
newbarrier
#()
{{
v
,
∃
l
:
loc
,
v
=
#
l
★
recv
l
P
★
send
l
P
}})
∧
(
∀
l
P
,
{{
send
l
P
★
P
}}
signal
#
l
{{
_
,
True
}})
∧
(
∀
l
P
,
{{
recv
l
P
}}
wait
#
l
{{
_
,
P
}})
∧
(
∀
l
P
Q
,
recv
l
(
P
★
Q
)
={
N
}=>
recv
l
P
★
recv
l
Q
)
∧
(
∀
l
P
Q
,
(
P
-
★
Q
)
⊢
recv
l
P
-
★
recv
l
Q
).
Proof
.
intros
HN
.
exists
(
λ
l
,
CofeMor
(
recv
heapN
N
l
)),
(
λ
l
,
CofeMor
(
send
heapN
N
l
)).
exists
(
λ
l
,
CofeMor
(
recv
N
l
)),
(
λ
l
,
CofeMor
(
send
N
l
)).
split_and
?
;
simpl
.
-
iIntros
(
P
)
"#? ! _"
.
iApply
(
newbarrier_spec
_
_
P
)
;
eauto
.
-
iIntros
(
P
)
"#? ! _"
.
iApply
(
newbarrier_spec
_
P
)
;
eauto
.
-
iIntros
(
l
P
)
"! [Hl HP]"
.
by
iApply
signal_spec
;
iFrame
"Hl HP"
.
-
iIntros
(
l
P
)
"! Hl"
.
iApply
wait_spec
;
iFrame
"Hl"
;
eauto
.
-
intros
;
by
apply
recv_split
.
...
...
heap_lang/lib/counter.v
View file @
6e487530
...
...
@@ -20,14 +20,13 @@ Instance inGF_counterG `{H : inGFs heap_lang Σ counterGF} : counterG Σ.
Proof
.
destruct
H
;
split
;
apply
_
.
Qed
.
Section
proof
.
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
counterG
Σ
}.
Context
(
heapN
:
namespace
).
Context
`
{!
heapG
Σ
,
!
counterG
Σ
}.
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Definition
counter_inv
(
l
:
loc
)
(
n
:
mnat
)
:
iProp
:
=
(
l
↦
#
n
)%
I
.
Definition
counter
(
l
:
loc
)
(
n
:
nat
)
:
iProp
:
=
(
∃
N
γ
,
heapN
⊥
N
∧
heap_ctx
heapN
∧
(
∃
N
γ
,
heapN
⊥
N
∧
heap_ctx
∧
auth_ctx
γ
N
(
counter_inv
l
)
∧
auth_own
γ
(
n
:
mnat
))%
I
.
(** The main proofs. *)
...
...
@@ -36,7 +35,7 @@ Proof. apply _. Qed.
Lemma
newcounter_spec
N
(
R
:
iProp
)
Φ
:
heapN
⊥
N
→
heap_ctx
heapN
★
(
∀
l
,
counter
l
0
-
★
Φ
#
l
)
⊢
WP
newcounter
#()
{{
Φ
}}.
heap_ctx
★
(
∀
l
,
counter
l
0
-
★
Φ
#
l
)
⊢
WP
newcounter
#()
{{
Φ
}}.
Proof
.
iIntros
(?)
"[#Hh HΦ]"
.
rewrite
/
newcounter
.
wp_seq
.
wp_alloc
l
as
"Hl"
.
iPvs
(
auth_alloc
(
counter_inv
l
)
N
_
(
O
:
mnat
)
with
"[Hl]"
)
...
...
heap_lang/lib/lock.v
View file @
6e487530
...
...
@@ -18,18 +18,17 @@ Instance inGF_lockG `{H : inGFs heap_lang Σ lockGF} : lockG Σ.
Proof
.
destruct
H
.
split
.
apply
:
inGF_inG
.
Qed
.
Section
proof
.
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
lockG
Σ
}.
Context
(
heapN
:
namespace
).
Context
`
{!
heapG
Σ
,
!
lockG
Σ
}.
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Definition
lock_inv
(
γ
:
gname
)
(
l
:
loc
)
(
R
:
iProp
)
:
iProp
:
=
(
∃
b
:
bool
,
l
↦
#
b
★
if
b
then
True
else
own
γ
(
Excl
())
★
R
)%
I
.
Definition
is_lock
(
l
:
loc
)
(
R
:
iProp
)
:
iProp
:
=
(
∃
N
γ
,
heapN
⊥
N
∧
heap_ctx
heapN
∧
inv
N
(
lock_inv
γ
l
R
))%
I
.
(
∃
N
γ
,
heapN
⊥
N
∧
heap_ctx
∧
inv
N
(
lock_inv
γ
l
R
))%
I
.
Definition
locked
(
l
:
loc
)
(
R
:
iProp
)
:
iProp
:
=
(
∃
N
γ
,
heapN
⊥
N
∧
heap_ctx
heapN
∧
(
∃
N
γ
,
heapN
⊥
N
∧
heap_ctx
∧
inv
N
(
lock_inv
γ
l
R
)
∧
own
γ
(
Excl
()))%
I
.
Global
Instance
lock_inv_ne
n
γ
l
:
Proper
(
dist
n
==>
dist
n
)
(
lock_inv
γ
l
).
...
...
@@ -48,7 +47,7 @@ Proof. rewrite /is_lock. iDestruct 1 as (N γ) "(?&?&?&_)"; eauto. Qed.
Lemma
newlock_spec
N
(
R
:
iProp
)
Φ
:
heapN
⊥
N
→
heap_ctx
heapN
★
R
★
(
∀
l
,
is_lock
l
R
-
★
Φ
#
l
)
⊢
WP
newlock
#()
{{
Φ
}}.
heap_ctx
★
R
★
(
∀
l
,
is_lock
l
R
-
★
Φ
#
l
)
⊢
WP
newlock
#()
{{
Φ
}}.
Proof
.
iIntros
(?)
"(#Hh & HR & HΦ)"
.
rewrite
/
newlock
.
wp_seq
.
wp_alloc
l
as
"Hl"
.
...
...
heap_lang/lib/par.v
View file @
6e487530
...
...
@@ -12,13 +12,12 @@ Notation "e1 || e2" := (par (Pair (λ: <>, e1) (λ: <>, e2)))%E : expr_scope.
Global
Opaque
par
.
Section
proof
.
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
spawnG
Σ
}.
Context
(
heapN
N
:
namespace
).
Context
`
{!
heapG
Σ
,
!
spawnG
Σ
}
(
N
:
namespace
).
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Lemma
par_spec
(
Ψ
1
Ψ
2
:
val
→
iProp
)
e
(
f1
f2
:
val
)
(
Φ
:
val
→
iProp
)
:
heapN
⊥
N
→
to_val
e
=
Some
(
f1
,
f2
)%
V
→
(
heap_ctx
heapN
★
WP
f1
#()
{{
Ψ
1
}}
★
WP
f2
#()
{{
Ψ
2
}}
★
(
heap_ctx
★
WP
f1
#()
{{
Ψ
1
}}
★
WP
f2
#()
{{
Ψ
2
}}
★
∀
v1
v2
,
Ψ
1
v1
★
Ψ
2
v2
-
★
▷
Φ
(
v1
,
v2
)%
V
)
⊢
WP
par
e
{{
Φ
}}.
Proof
.
...
...
@@ -34,7 +33,7 @@ Qed.
Lemma
wp_par
(
Ψ
1
Ψ
2
:
val
→
iProp
)
(
e1
e2
:
expr
)
`
{!
Closed
[]
e1
,
Closed
[]
e2
}
(
Φ
:
val
→
iProp
)
:
heapN
⊥
N
→
(
heap_ctx
heapN
★
WP
e1
{{
Ψ
1
}}
★
WP
e2
{{
Ψ
2
}}
★
(
heap_ctx
★
WP
e1
{{
Ψ
1
}}
★
WP
e2
{{
Ψ
2
}}
★
∀
v1
v2
,
Ψ
1
v1
★
Ψ
2
v2
-
★
▷
Φ
(
v1
,
v2
)%
V
)
⊢
WP
e1
||
e2
{{
Φ
}}.
Proof
.
...
...
heap_lang/lib/spawn.v
View file @
6e487530
...
...
@@ -28,8 +28,7 @@ Proof. destruct H as (?&?). split. apply: inGF_inG. Qed.
(** Now we come to the Iris part of the proof. *)
Section
proof
.
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
spawnG
Σ
}.
Context
(
heapN
N
:
namespace
).
Context
`
{!
heapG
Σ
,
!
spawnG
Σ
}
(
N
:
namespace
).
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Definition
spawn_inv
(
γ
:
gname
)
(
l
:
loc
)
(
Ψ
:
val
→
iProp
)
:
iProp
:
=
...
...
@@ -37,7 +36,7 @@ Definition spawn_inv (γ : gname) (l : loc) (Ψ : val → iProp) : iProp :=
∃
v
,
lv
=
SOMEV
v
★
(
Ψ
v
∨
own
γ
(
Excl
()))))%
I
.
Definition
join_handle
(
l
:
loc
)
(
Ψ
:
val
→
iProp
)
:
iProp
:
=
(
heapN
⊥
N
★
∃
γ
,
heap_ctx
heapN
★
own
γ
(
Excl
())
★
(
heapN
⊥
N
★
∃
γ
,
heap_ctx
★
own
γ
(
Excl
())
★
inv
N
(
spawn_inv
γ
l
Ψ
))%
I
.
Typeclasses
Opaque
join_handle
.
...
...
@@ -53,7 +52,7 @@ Proof. solve_proper. Qed.
Lemma
spawn_spec
(
Ψ
:
val
→
iProp
)
e
(
f
:
val
)
(
Φ
:
val
→
iProp
)
:
to_val
e
=
Some
f
→
heapN
⊥
N
→
heap_ctx
heapN
★
WP
f
#()
{{
Ψ
}}
★
(
∀
l
,
join_handle
l
Ψ
-
★
Φ
#
l
)
heap_ctx
★
WP
f
#()
{{
Ψ
}}
★
(
∀
l
,
join_handle
l
Ψ
-
★
Φ
#
l
)
⊢
WP
spawn
e
{{
Φ
}}.
Proof
.
iIntros
(<-%
of_to_val
?)
"(#Hh & Hf & HΦ)"
.
rewrite
/
spawn
.
...
...
heap_lang/proofmode.v
View file @
6e487530
...
...
@@ -6,8 +6,7 @@ Import uPred.
Ltac
wp_strip_later
::
=
iNext
.
Section
heap
.
Context
{
Σ
:
gFunctors
}
`
{
heapG
Σ
}.
Implicit
Types
N
:
namespace
.
Context
`
{
heapG
Σ
}.
Implicit
Types
P
Q
:
iPropG
heap_lang
Σ
.
Implicit
Types
Φ
:
val
→
iPropG
heap_lang
Σ
.
Implicit
Types
Δ
:
envs
(
iResUR
heap_lang
(
globalF
Σ
)).
...
...
@@ -16,9 +15,9 @@ Global Instance into_sep_mapsto l q v :
IntoSep
false
(
l
↦
{
q
}
v
)
(
l
↦
{
q
/
2
}
v
)
(
l
↦
{
q
/
2
}
v
).
Proof
.
by
rewrite
/
IntoSep
heap_mapsto_op_split
.
Qed
.
Lemma
tac_wp_alloc
Δ
Δ
'
N
E
j
e
v
Φ
:
Lemma
tac_wp_alloc
Δ
Δ
'
E
j
e
v
Φ
:
to_val
e
=
Some
v
→
(
Δ
⊢
heap_ctx
N
)
→
nclose
N
⊆
E
→
(
Δ
⊢
heap_ctx
)
→
nclose
heap
N
⊆
E
→
IntoLaterEnvs
Δ
Δ
'
→
(
∀
l
,
∃
Δ
''
,
envs_app
false
(
Esnoc
Enil
j
(
l
↦
v
))
Δ
'
=
Some
Δ
''
∧
...
...
@@ -32,8 +31,8 @@ Proof.
by
rewrite
right_id
H
Δ
'
.
Qed
.
Lemma
tac_wp_load
Δ
Δ
'
N
E
i
l
q
v
Φ
:
(
Δ
⊢
heap_ctx
N
)
→
nclose
N
⊆
E
→
Lemma
tac_wp_load
Δ
Δ
'
E
i
l
q
v
Φ
:
(
Δ
⊢
heap_ctx
)
→
nclose
heap
N
⊆
E
→
IntoLaterEnvs
Δ
Δ
'
→
envs_lookup
i
Δ
'
=
Some
(
false
,
l
↦
{
q
}
v
)%
I
→
(
Δ
'
⊢
|={
E
}=>
Φ
v
)
→
...
...
@@ -44,9 +43,9 @@ Proof.
by
apply
later_mono
,
sep_mono_r
,
wand_mono
.
Qed
.
Lemma
tac_wp_store
Δ
Δ
'
Δ
''
N
E
i
l
v
e
v'
Φ
:
Lemma
tac_wp_store
Δ
Δ
'
Δ
''
E
i
l
v
e
v'
Φ
:
to_val
e
=
Some
v'
→
(
Δ
⊢
heap_ctx
N
)
→
nclose
N
⊆
E
→
(
Δ
⊢
heap_ctx
)
→
nclose
heap
N
⊆
E
→
IntoLaterEnvs
Δ
Δ
'
→
envs_lookup
i
Δ
'
=
Some
(
false
,
l
↦
v
)%
I
→
envs_simple_replace
i
false
(
Esnoc
Enil
i
(
l
↦
v'
))
Δ
'
=
Some
Δ
''
→
...
...
@@ -58,9 +57,9 @@ Proof.
rewrite
right_id
.
by
apply
later_mono
,
sep_mono_r
,
wand_mono
.
Qed
.
Lemma
tac_wp_cas_fail
Δ
Δ
'
N
E
i
l
q
v
e1
v1
e2
v2
Φ
:
Lemma
tac_wp_cas_fail
Δ
Δ
'
E
i
l
q
v
e1
v1
e2
v2
Φ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
(
Δ
⊢
heap_ctx
N
)
→
nclose
N
⊆
E
→
(
Δ
⊢
heap_ctx
)
→
nclose
heap
N
⊆
E
→
IntoLaterEnvs
Δ
Δ
'
→
envs_lookup
i
Δ
'
=
Some
(
false
,
l
↦
{
q
}
v
)%
I
→
v
≠
v1
→
(
Δ
'
⊢
|={
E
}=>
Φ
(
LitV
(
LitBool
false
)))
→
...
...
@@ -71,9 +70,9 @@ Proof.
by
apply
later_mono
,
sep_mono_r
,
wand_mono
.
Qed
.
Lemma
tac_wp_cas_suc
Δ
Δ
'
Δ
''
N
E
i
l
v
e1
v1
e2
v2
Φ
:
Lemma
tac_wp_cas_suc
Δ
Δ
'
Δ
''
E
i
l
v
e1
v1
e2
v2
Φ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
(
Δ
⊢
heap_ctx
N
)
→
nclose
N
⊆
E
→
(
Δ
⊢
heap_ctx
)
→
nclose
heap
N
⊆
E
→
IntoLaterEnvs
Δ
Δ
'
→
envs_lookup
i
Δ
'
=
Some
(
false
,
l
↦
v
)%
I
→
v
=
v1
→
envs_simple_replace
i
false
(
Esnoc
Enil
i
(
l
↦
v2
))
Δ
'
=
Some
Δ
''
→
...
...
@@ -101,7 +100,7 @@ Tactic Notation "wp_alloc" ident(l) "as" constr(H) :=
[
reshape_expr
e
ltac
:
(
fun
K
e'
=>
match
eval
hnf
in
e'
with
Alloc
_
=>
wp_bind
K
end
)
|
fail
1
"wp_alloc: cannot find 'Alloc' in"
e
]
;
eapply
tac_wp_alloc
with
_
_
H
_;
eapply
tac_wp_alloc
with
_
H
_;
[
let
e'
:
=
match
goal
with
|-
to_val
?e'
=
_
=>
e'
end
in
wp_done
||
fail
"wp_alloc:"
e'
"not a value"
|
iAssumption
||
fail
"wp_alloc: cannot find heap_ctx"
...
...
tests/barrier_client.v
View file @
6e487530
...
...
@@ -14,7 +14,7 @@ Definition client : expr :=
Global
Opaque
worker
client
.
Section
client
.
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
barrierG
Σ
,
!
spawnG
Σ
}
(
heapN
N
:
namespace
).
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
barrierG
Σ
,
!
spawnG
Σ
}
(
N
:
namespace
).
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Definition
y_inv
(
q
:
Qp
)
(
l
:
loc
)
:
iProp
:
=
...
...
@@ -27,8 +27,7 @@ Section client.
Qed
.
Lemma
worker_safe
q
(
n
:
Z
)
(
b
y
:
loc
)
:
heap_ctx
heapN
★
recv
heapN
N
b
(
y_inv
q
y
)
⊢
WP
worker
n
#
b
#
y
{{
_
,
True
}}.
heap_ctx
★
recv
N
b
(
y_inv
q
y
)
⊢
WP
worker
n
#
b
#
y
{{
_
,
True
}}.
Proof
.
iIntros
"[#Hh Hrecv]"
.
wp_lam
.
wp_let
.
wp_apply
wait_spec
;
iFrame
"Hrecv"
.
...
...
@@ -37,12 +36,12 @@ Section client.
iApply
wp_wand_r
;
iSplitR
;
[
iApply
"Hf"
|
by
iIntros
(
v
)
"_"
].
Qed
.
Lemma
client_safe
:
heapN
⊥
N
→
heap_ctx
heapN
⊢
WP
client
{{
_
,
True
}}.
Lemma
client_safe
:
heapN
⊥
N
→
heap_ctx
⊢
WP
client
{{
_
,
True
}}.
Proof
.
iIntros
(?)
"#Hh"
;
rewrite
/
client
.
wp_alloc
y
as
"Hy"
.
wp_let
.
wp_apply
(
newbarrier_spec
heapN
N
(
y_inv
1
y
))
;
first
done
.
wp_apply
(
newbarrier_spec
N
(
y_inv
1
y
))
;
first
done
.
iFrame
"Hh"
.
iIntros
(
l
)
"[Hr Hs]"
.
wp_let
.
iApply
(
wp_par
heapN
N
(
λ
_
,
True
%
I
)
(
λ
_
,
True
%
I
))
;
first
done
.
iApply
(
wp_par
N
(
λ
_
,
True
%
I
)
(
λ
_
,
True
%
I
))
;
first
done
.
iFrame
"Hh"
.
iSplitL
"Hy Hs"
.
-
(* The original thread, the sender. *)
wp_store
.
iApply
signal_spec
;
iFrame
"Hs"
;
iSplit
;
[|
done
].
...
...
@@ -52,7 +51,7 @@ Section client.
iDestruct
(
recv_weaken
with
"[] Hr"
)
as
"Hr"
.
{
iIntros
"Hy"
.
by
iApply
(
y_inv_split
with
"Hy"
).
}
iPvs
(
recv_split
with
"Hr"
)
as
"[H1 H2]"
;
first
done
.
iApply
(
wp_par
heapN
N
(
λ
_
,
True
%
I
)
(
λ
_
,
True
%
I
))
;
eauto
.
iApply
(
wp_par
N
(
λ
_
,
True
%
I
)
(
λ
_
,
True
%
I
))
;
eauto
.
iFrame
"Hh"
;
iSplitL
"H1"
;
[|
iSplitL
"H2"
;
[|
by
iIntros
(
_
_
)
"_ >"
]]
;
iApply
worker_safe
;
by
iSplit
.
Qed
.
...
...
@@ -65,8 +64,8 @@ Section ClosedProofs.
Lemma
client_safe_closed
σ
:
{{
ownP
σ
:
iProp
}}
client
{{
v
,
True
}}.
Proof
.
iIntros
"! Hσ"
.
iPvs
(
heap_alloc
(
nroot
.@
"Barrier"
)
with
"Hσ"
)
as
(
h
)
"[#Hh _]"
;
first
done
.
iApply
(
client_safe
(
nroot
.@
"
B
arrier"
)
(
nroot
.@
"Heap"
)
)
;
auto
with
ndisj
.
iPvs
(
heap_alloc
with
"Hσ"
)
as
(
h
)
"[#Hh _]"
;
first
done
.
iApply
(
client_safe
(
nroot
.@
"
b
arrier"
))
;
auto
with
ndisj
.
Qed
.
Print
Assumptions
client_safe_closed
.
...
...
tests/heap_lang.v
View file @
6e487530
...
...
@@ -23,21 +23,21 @@ Section LiftingTests.
Definition
heap_e
:
expr
:
=
let
:
"x"
:
=
ref
#
1
in
"x"
<-
!
"x"
+
#
1
;;
!
"x"
.
Lemma
heap_e_spec
E
N
:
nclose
N
⊆
E
→
heap_ctx
N
⊢
WP
heap_e
@
E
{{
v
,
v
=
#
2
}}.
Lemma
heap_e_spec
E
:
nclose
heap
N
⊆
E
→
heap_ctx
⊢
WP
heap_e
@
E
{{
v
,
v
=
#
2
}}.
Proof
.
iIntros
(
HN
)
"#?"
.
rewrite
/
heap_e
.
iApply
(
wp_mask_weaken
N
)
;
first
done
.
iIntros
(
HN
)
"#?"
.
rewrite
/
heap_e
.
wp_alloc
l
.
wp_let
.
wp_load
.
wp_op
.
wp_store
.
by
wp_load
.
Qed
.
Definition
heap_e2
:
expr
:
=
Definition
heap_e2
:
expr
:
=
let
:
"x"
:
=
ref
#
1
in
let
:
"y"
:
=
ref
#
1
in
"x"
<-
!
"x"
+
#
1
;;
!
"x"
.
Lemma
heap_e2_spec
E
N
:
nclose
N
⊆
E
→
heap_ctx
N
⊢
WP
heap_e2
@
E
{{
v
,
v
=
#
2
}}.
Lemma
heap_e2_spec
E
:
nclose
heap
N
⊆
E
→
heap_ctx
⊢
WP
heap_e2
@
E
{{
v
,
v
=
#
2
}}.
Proof
.
iIntros
(
HN
)
"#?"
.
rewrite
/
heap_e2
.
iApply
(
wp_mask_weaken
N
)
;
first
done
.
iIntros
(
HN
)
"#?"
.
rewrite
/
heap_e2
.
wp_alloc
l
.
wp_let
.
wp_alloc
l'
.
wp_let
.
wp_load
.
wp_op
.
wp_store
.
wp_load
.
done
.
Qed
.
...
...
@@ -82,7 +82,7 @@ Section ClosedProofs.
Lemma
heap_e_closed
σ
:
{{
ownP
σ
:
iProp
}}
heap_e
{{
v
,
v
=
#
2
}}.
Proof
.
iProof
.
iIntros
"! Hσ"
.
iPvs
(
heap_alloc
nroot
with
"Hσ"
)
as
(
h
)
"[? _]"
;
first
by
rewrite
nclose_nroot
.
iApply
heap_e_spec
;
last
done
;
by
rewrite
nclose_nroot
.
iPvs
(
heap_alloc
with
"Hσ"
)
as
(
h
)
"[? _]"
;
first
solve_ndisj
.
by
iApply
heap_e_spec
;
first
solve_ndisj
.
Qed
.
End
ClosedProofs
.
tests/joining_existentials.v
View file @
6e487530
...
...
@@ -20,8 +20,8 @@ Global Opaque client.
Section
proof
.
Context
(
G
:
cFunctor
).
Context
{
Σ
:
gFunctors
}
`
{!
heapG
Σ
,
!
barrierG
Σ
,
!
spawnG
Σ
,
!
oneShotG
heap_lang
Σ
G
}.
Context
(
heapN
N
:
namespace
).
Context
`
{!
heapG
Σ
,
!
barrierG
Σ
,
!
spawnG
Σ
,
!
oneShotG
heap_lang
Σ
G
}.
Context
(
N
:
namespace
).
Local
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Local
Notation
X
:
=
(
G
iProp
).
...
...
@@ -30,7 +30,7 @@ Definition barrier_res γ (Φ : X → iProp) : iProp :=
Next
(
cFunctor_map
G
(
iProp_fold
,
iProp_unfold
)
x
))
★
Φ
x
)%
I
.
Lemma
worker_spec
e
γ
l
(
Φ
Ψ
:
X
→
iProp
)
`
{!
Closed
[]
e
}
:
recv
heapN
N
l
(
barrier_res
γ
Φ
)
★
(
∀
x
,
{{
Φ
x
}}
e
{{
_
,
Ψ
x
}})
recv
N
l
(
barrier_res
γ
Φ
)
★
(
∀
x
,
{{
Φ
x
}}
e
{{
_
,
Ψ
x
}})
⊢
WP
wait
#
l
;;
e
{{
_
,
barrier_res
γ
Ψ
}}.
Proof
.
iIntros
"[Hl #He]"
.
wp_apply
wait_spec
;
iFrame
"Hl"
.
...
...
@@ -66,7 +66,7 @@ Qed.
Lemma
client_spec_new
eM
eW1
eW2
`
{!
Closed
[]
eM
,
!
Closed
[]
eW1
,
!
Closed
[]
eW2
}
:
heapN
⊥
N
→
heap_ctx
heapN
★
P
heap_ctx
★
P
★
{{
P
}}
eM
{{
_
,
∃
x
,
Φ
x
}}
★
(
∀
x
,
{{
Φ
1
x
}}
eW1
{{
_
,
Ψ
1
x
}})
★
(
∀
x
,
{{
Φ
2
x
}}
eW2
{{
_
,
Ψ
2
x
}})
...
...
@@ -74,10 +74,10 @@ Lemma client_spec_new eM eW1 eW2 `{!Closed [] eM, !Closed [] eW1, !Closed [] eW2
Proof
.
iIntros
(
HN
)
"/= (#Hh&HP&#He&#He1&#He2)"
;
rewrite
/
client
.
iPvs
(
own_alloc
(
Cinl
(
Excl
())))
as
(
γ
)
"Hγ"
.
done
.
wp_apply
(
newbarrier_spec
heapN
N
(
barrier_res
γ
Φ
))
;
auto
.
wp_apply
(
newbarrier_spec
N
(
barrier_res
γ
Φ
))
;
auto
.
iFrame
"Hh"
.
iIntros
(
l
)
"[Hr Hs]"
.
set
(
workers_post
(
v
:
val
)
:
=
(
barrier_res
γ
Ψ
1
★
barrier_res
γ
Ψ
2
)%
I
).
wp_let
.
wp_apply
(
wp_par
_
_
(
λ
_
,
True
)%
I
workers_post
)
;
wp_let
.
wp_apply
(
wp_par
_
(
λ
_
,
True
)%
I
workers_post
)
;
try
iFrame
"Hh"
;
first
done
.
iSplitL
"HP Hs Hγ"
;
[|
iSplitL
"Hr"
].
-
wp_focus
eM
.
iApply
wp_wand_l
;
iSplitR
"HP"
;
[|
by
iApply
"He"
].
...
...
@@ -88,7 +88,7 @@ Proof.
iExists
x
;
auto
.
-
iDestruct
(
recv_weaken
with
"[] Hr"
)
as
"Hr"
;
first
by
iApply
P_res_split
.
iPvs
(
recv_split
with
"Hr"
)
as
"[H1 H2]"
;
first
done
.
wp_apply
(
wp_par
_
_
(
λ
_
,
barrier_res
γ
Ψ
1
)%
I
wp_apply
(
wp_par
_
(
λ
_
,
barrier_res
γ
Ψ
1
)%
I
(
λ
_
,
barrier_res
γ
Ψ
2
)%
I
)
;
try
iFrame
"Hh"
;
first
done
.
iSplitL
"H1"
;
[|
iSplitL
"H2"
].
+
iApply
worker_spec
;
auto
.
...
...
tests/list_reverse.v
View file @
6e487530
...
...
@@ -4,7 +4,7 @@ From iris.program_logic Require Export hoare.
From
iris
.
heap_lang
Require
Import
proofmode
notation
.
Section
list_reverse
.
Context
`
{!
heapG
Σ
}
(
heapN
:
namespace
)
.
Context
`
{!
heapG
Σ
}.
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Implicit
Types
l
:
loc
.
...
...
@@ -27,7 +27,7 @@ Definition rev : val :=
Global
Opaque
rev
.
Lemma
rev_acc_wp
hd
acc
xs
ys
(
Φ
:
val
→
iProp
)
:
heap_ctx
heapN
★
is_list
hd
xs
★
is_list
acc
ys
★
heap_ctx
★
is_list
hd
xs
★
is_list
acc
ys
★
(
∀
w
,
is_list
w
(
reverse
xs
++
ys
)
-
★
Φ
w
)
⊢
WP
rev
hd
acc
{{
Φ
}}.
Proof
.
...
...
@@ -43,7 +43,7 @@ Proof.
Qed
.
Lemma
rev_wp
hd
xs
(
Φ
:
val
→
iProp
)
:
heap_ctx
heapN
★
is_list
hd
xs
★
(
∀
w
,
is_list
w
(
reverse
xs
)
-
★
Φ
w
)
heap_ctx
★
is_list
hd
xs
★
(
∀
w
,
is_list
w
(
reverse
xs
)
-
★
Φ
w
)
⊢
WP
rev
hd
(
InjL
#())
{{
Φ
}}.
Proof
.
iIntros
"(#Hh & Hxs & HΦ)"
.
...
...