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
Rice Wine
Iris
Commits
75dd9e14
Commit
75dd9e14
authored
Oct 04, 2018
by
Ralf Jung
Browse files
Merge branch 'ralf/curry-fork' into 'master'
curry wp_fork See merge request FP/iris-coq!180
parents
f5109e7c
9d80e031
Changes
2
Hide whitespace changes
Inline
Side-by-side
theories/heap_lang/lib/spawn.v
View file @
75dd9e14
...
...
@@ -53,11 +53,11 @@ Proof.
iMod
(
own_alloc
(
Excl
()))
as
(
γ
)
"Hγ"
;
first
done
.
iMod
(
inv_alloc
N
_
(
spawn_inv
γ
l
Ψ
)
with
"[Hl]"
)
as
"#?"
.
{
iNext
.
iExists
NONEV
.
iFrame
;
eauto
.
}
wp_apply
wp_fork
;
simpl
.
iSplitR
"Hf"
.
-
wp_seq
.
iApply
"HΦ"
.
rewrite
/
join_handle
.
eauto
.
-
wp_bind
(
f
_
).
iApply
(
wp_wand
with
"Hf"
)
;
iIntros
(
v
)
"Hv"
.
wp_apply
(
wp_fork
with
"[Hf]"
).
-
iNext
.
wp_bind
(
f
_
).
iApply
(
wp_wand
with
"Hf"
)
;
iIntros
(
v
)
"Hv"
.
iInv
N
as
(
v'
)
"[Hl _]"
.
wp_store
.
iSplitL
;
last
done
.
iIntros
"!> !>"
.
iExists
(
SOMEV
v
).
iFrame
.
eauto
.
-
wp_seq
.
iApply
"HΦ"
.
rewrite
/
join_handle
.
eauto
.
Qed
.
Lemma
join_spec
(
Ψ
:
val
→
iProp
Σ
)
l
:
...
...
theories/heap_lang/lifting.v
View file @
75dd9e14
...
...
@@ -104,18 +104,18 @@ Implicit Types Φ : val → iProp Σ.
Implicit
Types
efs
:
list
expr
.
Implicit
Types
σ
:
state
.
(**
Base axioms for core primitives of the language: Stateless reductions
*)
(**
Fork: Not using Texan triples to avoid some unnecessary [True]
*)
Lemma
wp_fork
s
E
e
Φ
:
▷
(
Φ
(
LitV
LitUnit
)
∗
WP
e
@
s
;
⊤
{{
_
,
True
}}
)
⊢
WP
Fork
e
@
s
;
E
{{
Φ
}}.
▷
WP
e
@
s
;
⊤
{{
_
,
True
}}
-
∗
▷
Φ
(
LitV
LitUnit
)
-
∗
WP
Fork
e
@
s
;
E
{{
Φ
}}.
Proof
.
iIntros
"
[HΦ He]
"
.
iIntros
"
He HΦ
"
.
iApply
wp_lift_pure_det_head_step
;
[
auto
|
intros
;
inv_head_step
;
eauto
|].
iModIntro
;
iNext
;
iIntros
"!> /= {$He}"
.
by
iApply
wp_value
.
Qed
.
Lemma
twp_fork
s
E
e
Φ
:
Φ
(
LitV
LitUnit
)
∗
WP
e
@
s
;
⊤
[{
_
,
True
}]
⊢
WP
Fork
e
@
s
;
E
[{
Φ
}].
WP
e
@
s
;
⊤
[{
_
,
True
}]
-
∗
Φ
(
LitV
LitUnit
)
-
∗
WP
Fork
e
@
s
;
E
[{
Φ
}].
Proof
.
iIntros
"
[HΦ He]
"
.
iIntros
"
He HΦ
"
.
iApply
twp_lift_pure_det_head_step
;
[
auto
|
intros
;
inv_head_step
;
eauto
|].
iIntros
"!> /= {$He}"
.
by
iApply
twp_value
.
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