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
Rice Wine
Iris
Commits
4ec14182
Commit
4ec14182
authored
Oct 04, 2018
by
Ralf Jung
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
tweak WP def.n and fix eauto in heap_lang
parent
f7991ef5
Changes
8
Show whitespace changes
Inline
Side-by-side
Showing
8 changed files
with
124 additions
and
145 deletions
+124
-145
theories/heap_lang/lang.v
theories/heap_lang/lang.v
+1
-0
theories/heap_lang/lifting.v
theories/heap_lang/lifting.v
+33
-56
theories/program_logic/adequacy.v
theories/program_logic/adequacy.v
+3
-3
theories/program_logic/ectx_lifting.v
theories/program_logic/ectx_lifting.v
+33
-33
theories/program_logic/lifting.v
theories/program_logic/lifting.v
+24
-24
theories/program_logic/ownp.v
theories/program_logic/ownp.v
+8
-7
theories/program_logic/total_weakestpre.v
theories/program_logic/total_weakestpre.v
+2
-2
theories/program_logic/weakestpre.v
theories/program_logic/weakestpre.v
+20
-20
No files found.
theories/heap_lang/lang.v
View file @
4ec14182
...
@@ -167,6 +167,7 @@ Definition val_is_unboxed (v : val) : Prop :=
...
@@ -167,6 +167,7 @@ Definition val_is_unboxed (v : val) : Prop :=
(** The state: heaps of vals. *)
(** The state: heaps of vals. *)
Definition
state
:
Type
:
=
gmap
loc
val
*
gset
proph
.
Definition
state
:
Type
:
=
gmap
loc
val
*
gset
proph
.
Implicit
Type
σ
:
state
.
(** Equality and other typeclass stuff *)
(** Equality and other typeclass stuff *)
Lemma
to_of_val
v
:
to_val
(
of_val
v
)
=
Some
v
.
Lemma
to_of_val
v
:
to_val
(
of_val
v
)
=
Some
v
.
...
...
theories/heap_lang/lifting.v
View file @
4ec14182
...
@@ -48,11 +48,15 @@ Ltac inv_head_step :=
...
@@ -48,11 +48,15 @@ Ltac inv_head_step :=
end
.
end
.
Local
Hint
Extern
0
(
atomic
_
_
)
=>
solve_atomic
.
Local
Hint
Extern
0
(
atomic
_
_
)
=>
solve_atomic
.
Local
Hint
Extern
0
(
head_reducible
_
_
)
=>
eexists
_
,
_
,
_;
simpl
.
Local
Hint
Extern
0
(
head_reducible
_
_
)
=>
eexists
_
,
_
,
_
,
_;
simpl
.
Local
Hint
Extern
0
(
head_reducible_no_obs
_
_
)
=>
eexists
_
,
_;
simpl
.
Local
Hint
Extern
0
(
head_reducible_no_obs
_
_
)
=>
eexists
_
,
_
,
_;
simpl
.
Local
Hint
Constructors
head_step
.
(* [simpl apply] is too stupid, so we need extern hints here. *)
Local
Hint
Resolve
alloc_fresh
.
Local
Hint
Extern
1
(
head_step
_
_
_
_
_
_
)
=>
econstructor
.
Local
Hint
Extern
0
(
head_step
(
CAS
_
_
_
)
_
_
_
_
_
)
=>
eapply
CasSucS
.
Local
Hint
Extern
0
(
head_step
(
CAS
_
_
_
)
_
_
_
_
_
)
=>
eapply
CasFailS
.
Local
Hint
Extern
0
(
head_step
(
Alloc
_
)
_
_
_
_
_
)
=>
apply
alloc_fresh
.
Local
Hint
Extern
0
(
head_step
NewProph
_
_
_
_
_
)
=>
apply
new_proph_fresh
.
Local
Hint
Resolve
to_of_val
.
Local
Hint
Resolve
to_of_val
.
Local
Ltac
solve_exec_safe
:
=
intros
;
subst
;
do
3
eexists
;
econstructor
;
eauto
.
Local
Ltac
solve_exec_safe
:
=
intros
;
subst
;
do
3
eexists
;
econstructor
;
eauto
.
...
@@ -134,10 +138,8 @@ Lemma wp_alloc s E e v :
...
@@ -134,10 +138,8 @@ Lemma wp_alloc s E e v :
{{{
True
}}}
Alloc
e
@
s
;
E
{{{
l
,
RET
LitV
(
LitLoc
l
)
;
l
↦
v
}}}.
{{{
True
}}}
Alloc
e
@
s
;
E
{{{
l
,
RET
LitV
(
LitLoc
l
)
;
l
↦
v
}}}.
Proof
.
Proof
.
iIntros
(<-
Φ
)
"_ HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(<-
Φ
)
"_ HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
;
iSplit
.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ Hκs] !>"
;
iSplit
;
first
by
eauto
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
{
iPureIntro
.
repeat
eexists
.
by
apply
alloc_fresh
.
}
iNext
;
iIntros
(
κ
κ
s'
v2
σ
2
efs
[
Hstep
->])
;
inv_head_step
.
iMod
(@
gen_heap_alloc
with
"Hσ"
)
as
"[Hσ Hl]"
;
first
done
.
iMod
(@
gen_heap_alloc
with
"Hσ"
)
as
"[Hσ Hl]"
;
first
done
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -146,9 +148,7 @@ Lemma twp_alloc s E e v :
...
@@ -146,9 +148,7 @@ Lemma twp_alloc s E e v :
[[{
True
}]]
Alloc
e
@
s
;
E
[[{
l
,
RET
LitV
(
LitLoc
l
)
;
l
↦
v
}]].
[[{
True
}]]
Alloc
e
@
s
;
E
[[{
l
,
RET
LitV
(
LitLoc
l
)
;
l
↦
v
}]].
Proof
.
Proof
.
iIntros
(<-
Φ
)
"_ HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(<-
Φ
)
"_ HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
;
iSplit
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
;
iSplit
;
first
by
eauto
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
by
apply
alloc_fresh
.
}
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iMod
(@
gen_heap_alloc
with
"Hσ"
)
as
"[Hσ Hl]"
;
first
done
.
iMod
(@
gen_heap_alloc
with
"Hσ"
)
as
"[Hσ Hl]"
;
first
done
.
iModIntro
;
iSplit
=>
//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
...
@@ -158,9 +158,8 @@ Lemma wp_load s E l q v :
...
@@ -158,9 +158,8 @@ Lemma wp_load s E l q v :
{{{
▷
l
↦
{
q
}
v
}}}
Load
(
Lit
(
LitLoc
l
))
@
s
;
E
{{{
RET
v
;
l
↦
{
q
}
v
}}}.
{{{
▷
l
↦
{
q
}
v
}}}
Load
(
Lit
(
LitLoc
l
))
@
s
;
E
{{{
RET
v
;
l
↦
{
q
}
v
}}}.
Proof
.
Proof
.
iIntros
(
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iNext
;
iIntros
(
κ
κ
s'
v2
σ
2
efs
[
Hstep
->])
;
inv_head_step
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
Lemma
twp_load
s
E
l
q
v
:
Lemma
twp_load
s
E
l
q
v
:
...
@@ -168,8 +167,7 @@ Lemma twp_load s E l q v :
...
@@ -168,8 +167,7 @@ Lemma twp_load s E l q v :
Proof
.
Proof
.
iIntros
(
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iSplit
;
first
by
eauto
.
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iModIntro
;
iSplit
=>
//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -179,11 +177,8 @@ Lemma wp_store s E l v' e v :
...
@@ -179,11 +177,8 @@ Lemma wp_store s E l v' e v :
Proof
.
Proof
.
iIntros
(<-
Φ
)
">Hl HΦ"
.
iIntros
(<-
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
constructor
;
eauto
.
}
iNext
;
iIntros
(
κ
κ
s'
v2
σ
2
efs
[
Hstep
->])
;
inv_head_step
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -194,10 +189,7 @@ Proof.
...
@@ -194,10 +189,7 @@ Proof.
iIntros
(<-
Φ
)
"Hl HΦ"
.
iIntros
(<-
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
.
iSplit
;
first
by
eauto
.
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
constructor
;
eauto
.
}
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -209,8 +201,8 @@ Lemma wp_cas_fail s E l q v' e1 v1 e2 :
...
@@ -209,8 +201,8 @@ Lemma wp_cas_fail s E l q v' e1 v1 e2 :
Proof
.
Proof
.
iIntros
(<-
[
v2
<-]
??
Φ
)
">Hl HΦ"
.
iIntros
(<-
[
v2
<-]
??
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
κ
κ
s'
v2'
σ
2
efs
[
Hstep
->]
)
;
inv_head_step
.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
Lemma
twp_cas_fail
s
E
l
q
v'
e1
v1
e2
:
Lemma
twp_cas_fail
s
E
l
q
v'
e1
v1
e2
:
...
@@ -232,11 +224,8 @@ Lemma wp_cas_suc s E l e1 v1 e2 v2 :
...
@@ -232,11 +224,8 @@ Lemma wp_cas_suc s E l e1 v1 e2 v2 :
Proof
.
Proof
.
iIntros
(<-
<-
?
Φ
)
">Hl HΦ"
.
iIntros
(<-
<-
?
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
by
econstructor
.
}
iNext
;
iIntros
(
κ
κ
s'
v2'
σ
2
efs
[
Hstep
->])
;
inv_head_step
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -248,10 +237,7 @@ Proof.
...
@@ -248,10 +237,7 @@ Proof.
iIntros
(<-
<-
?
Φ
)
"Hl HΦ"
.
iIntros
(<-
<-
?
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
.
iSplit
;
first
by
eauto
.
iIntros
(
κ
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
by
econstructor
.
}
iIntros
(
κ
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -263,11 +249,8 @@ Lemma wp_faa s E l i1 e2 i2 :
...
@@ -263,11 +249,8 @@ Lemma wp_faa s E l i1 e2 i2 :
Proof
.
Proof
.
iIntros
(<-
Φ
)
">Hl HΦ"
.
iIntros
(<-
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
by
constructor
.
}
iNext
;
iIntros
(
κ
κ
s'
v2'
σ
2
efs
[
Hstep
->])
;
inv_head_step
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -279,10 +262,7 @@ Proof.
...
@@ -279,10 +262,7 @@ Proof.
iIntros
(<-
Φ
)
"Hl HΦ"
.
iIntros
(<-
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
κ
s
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
.
iSplit
;
first
by
eauto
.
iIntros
(
κ
e2
σ
2
efs
Hstep
)
;
inv_head_step
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
by
constructor
.
}
iIntros
(
κ
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -292,12 +272,10 @@ Lemma wp_new_proph :
...
@@ -292,12 +272,10 @@ Lemma wp_new_proph :
{{{
True
}}}
NewProph
{{{
v
(
p
:
proph
),
RET
(
LitV
(
LitProphecy
p
))
;
p
⥱
v
}}}.
{{{
True
}}}
NewProph
{{{
v
(
p
:
proph
),
RET
(
LitV
(
LitProphecy
p
))
;
p
⥱
v
}}}.
Proof
.
Proof
.
iIntros
(
Φ
)
"_ HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
Φ
)
"_ HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ HR] !>"
.
iDestruct
"HR"
as
(
R
[
Hfr
Hdom
])
"HR"
.
iSplit
.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ HR] !>"
.
iDestruct
"HR"
as
(
R
[
Hfr
Hdom
])
"HR"
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
iSplit
;
first
by
eauto
.
{
iPureIntro
.
repeat
eexists
.
by
apply
new_proph_fresh
.
}
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
).
inv_head_step
.
unfold
cons_obs
.
simpl
.
iMod
(@
proph_map_alloc
with
"HR"
)
as
"[HR Hp]"
.
iNext
;
iIntros
(
κ
κ
s'
v2
σ
2
efs
[
Hstep
->]).
inv_head_step
.
iMod
((@
proph_map_alloc
_
_
_
_
_
_
_
p
)
with
"HR"
)
as
"[HR Hp]"
.
{
intro
Hin
.
apply
(
iffLR
(
elem_of_subseteq
_
_
)
Hdom
)
in
Hin
.
done
.
}
{
intro
Hin
.
apply
(
iffLR
(
elem_of_subseteq
_
_
)
Hdom
)
in
Hin
.
done
.
}
iModIntro
;
iSplit
=>
//.
iFrame
.
iSplitL
"HR"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
iSplitL
"HR"
.
-
iExists
_
.
iSplit
;
last
done
.
-
iExists
_
.
iSplit
;
last
done
.
...
@@ -313,12 +291,11 @@ Lemma wp_resolve_proph e1 e2 p v w:
...
@@ -313,12 +291,11 @@ Lemma wp_resolve_proph e1 e2 p v w:
{{{
p
⥱
v
}}}
ResolveProph
e1
e2
{{{
RET
(
LitV
LitUnit
)
;
⌜
v
=
Some
w
⌝
}}}.
{{{
p
⥱
v
}}}
ResolveProph
e1
e2
{{{
RET
(
LitV
LitUnit
)
;
⌜
v
=
Some
w
⌝
}}}.
Proof
.
Proof
.
iIntros
(<-
<-
Φ
)
"Hp HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(<-
<-
Φ
)
"Hp HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
)
"[Hσ HR] !>"
.
iDestruct
"HR"
as
(
R
[
Hfr
Hdom
])
"HR"
.
iDestruct
(@
proph_map_valid
with
"HR Hp"
)
as
%
Hlookup
.
iSplit
.
(* TODO (MR) this used to be done by eauto. Why does it not work any more? *)
{
iPureIntro
.
repeat
eexists
.
by
constructor
.
}
unfold
cons_obs
.
simpl
.
unfold
cons_obs
.
simpl
.
iNext
;
iIntros
(
κ
κ
s'
v2
σ
2
efs
[
Hstep
->])
;
inv_head_step
.
iApply
fupd_frame_l
.
iIntros
(
σ
1
κ
κ
s
)
"[Hσ HR] !>"
.
iDestruct
"HR"
as
(
R
[
Hfr
Hdom
])
"HR"
.
iDestruct
(@
proph_map_valid
with
"HR Hp"
)
as
%
Hlookup
.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iApply
fupd_frame_l
.
iSplit
=>
//.
iFrame
.
iSplit
=>
//.
iFrame
.
iMod
(@
proph_map_remove
with
"HR Hp"
)
as
"Hp"
.
iModIntro
.
iMod
(@
proph_map_remove
with
"HR Hp"
)
as
"Hp"
.
iModIntro
.
iSplitR
"HΦ"
.
iSplitR
"HΦ"
.
...
...
theories/program_logic/adequacy.v
View file @
4ec14182
...
@@ -78,8 +78,8 @@ Lemma wp_step s E e1 σ1 κ κs e2 σ2 efs Φ :
...
@@ -78,8 +78,8 @@ Lemma wp_step s E e1 σ1 κ κs e2 σ2 efs Φ :
Proof
.
Proof
.
rewrite
{
1
}
wp_unfold
/
wp_pre
.
iIntros
(?)
"[(Hw & HE & Hσ) H]"
.
rewrite
{
1
}
wp_unfold
/
wp_pre
.
iIntros
(?)
"[(Hw & HE & Hσ) H]"
.
rewrite
(
val_stuck
e1
σ
1
κ
e2
σ
2
efs
)
//
uPred_fupd_eq
.
rewrite
(
val_stuck
e1
σ
1
κ
e2
σ
2
efs
)
//
uPred_fupd_eq
.
iMod
(
"H"
$!
σ
1
_
with
"Hσ [Hw HE]"
)
as
">(Hw & HE & _ & H)"
;
first
by
iFrame
.
iMod
(
"H"
$!
σ
1
with
"Hσ [Hw HE]"
)
as
">(Hw & HE & _ & H)"
;
first
by
iFrame
.
iMod
(
"H"
$!
κ
κ
s
e2
σ
2
efs
with
"[//] [$Hw $HE]"
)
as
">(Hw & HE & H)"
.
iMod
(
"H"
$!
e2
σ
2
efs
with
"[//] [$Hw $HE]"
)
as
">(Hw & HE & H)"
.
iIntros
"!> !>"
.
by
iMod
(
"H"
with
"[$Hw $HE]"
)
as
">($ & $ & $)"
.
iIntros
"!> !>"
.
by
iMod
(
"H"
with
"[$Hw $HE]"
)
as
">($ & $ & $)"
.
Qed
.
Qed
.
...
@@ -145,7 +145,7 @@ Proof.
...
@@ -145,7 +145,7 @@ Proof.
rewrite
wp_unfold
/
wp_pre
.
iIntros
"(Hw&HE&Hσ) H"
.
rewrite
wp_unfold
/
wp_pre
.
iIntros
"(Hw&HE&Hσ) H"
.
destruct
(
to_val
e
)
as
[
v
|]
eqn
:
?.
destruct
(
to_val
e
)
as
[
v
|]
eqn
:
?.
{
iIntros
"!> !> !%"
.
left
.
by
exists
v
.
}
{
iIntros
"!> !> !%"
.
left
.
by
exists
v
.
}
rewrite
uPred_fupd_eq
.
iMod
(
"H"
with
"Hσ [-]"
)
as
">(?&?&%&?)"
;
first
by
iFrame
.
rewrite
uPred_fupd_eq
.
iMod
(
"H"
$!
_
None
with
"Hσ [-]"
)
as
">(?&?&%&?)"
;
first
by
iFrame
.
iIntros
"!> !> !%"
.
by
right
.
iIntros
"!> !> !%"
.
by
right
.
Qed
.
Qed
.
...
...
theories/program_logic/ectx_lifting.v
View file @
4ec14182
...
@@ -16,28 +16,28 @@ Hint Resolve head_stuck_stuck.
...
@@ -16,28 +16,28 @@ Hint Resolve head_stuck_stuck.
Lemma
wp_lift_head_step_fupd
{
s
E
Φ
}
e1
:
Lemma
wp_lift_head_step_fupd
{
s
E
Φ
}
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E
,
∅
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E
,
∅
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
⌜
head_reducible
e1
σ
1
⌝
∗
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
∅
,
∅
,
E
}
▷
=
∗
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
∅
,
∅
,
E
}
▷
=
∗
state_interp
σ
2
κ
s
'
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
state_interp
σ
2
κ
s
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
Proof
.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_step_fupd
=>//.
iIntros
(
σ
1
κ
s
)
"Hσ"
.
iIntros
(?)
"H"
.
iApply
wp_lift_step_fupd
=>//.
iIntros
(
σ
1
κ
κ
s
)
"Hσ"
.
iMod
(
"H"
with
"Hσ"
)
as
"[% H]"
;
iModIntro
.
iMod
(
"H"
with
"Hσ"
)
as
"[% H]"
;
iModIntro
.
iSplit
;
first
by
destruct
s
;
eauto
.
iIntros
(
κ
κ
s'
e2
σ
2
efs
[
Hstep
->]
).
iSplit
;
first
by
destruct
s
;
eauto
.
iIntros
(
e2
σ
2
efs
Hstep
).
iApply
"H"
;
eauto
.
iApply
"H"
;
eauto
.
Qed
.
Qed
.
Lemma
wp_lift_head_step
{
s
E
Φ
}
e1
:
Lemma
wp_lift_head_step
{
s
E
Φ
}
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E
,
∅
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E
,
∅
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
⌜
head_reducible
e1
σ
1
⌝
∗
▷
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
∅
,
E
}=
∗
▷
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
∅
,
E
}=
∗
state_interp
σ
2
κ
s
'
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
state_interp
σ
2
κ
s
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
Proof
.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_head_step_fupd
;
[
done
|].
iIntros
(??)
"?"
.
iIntros
(?)
"H"
.
iApply
wp_lift_head_step_fupd
;
[
done
|].
iIntros
(??
?
)
"?"
.
iMod
(
"H"
with
"[$]"
)
as
"[$ H]"
.
iIntros
"!>"
(
κ
κ
s'
e2
σ
2
efs
?)
"!> !>"
.
by
iApply
"H"
.
iMod
(
"H"
with
"[$]"
)
as
"[$ H]"
.
iIntros
"!>"
(
e2
σ
2
efs
?)
"!> !>"
.
by
iApply
"H"
.
Qed
.
Qed
.
Lemma
wp_lift_head_stuck
E
Φ
e
:
Lemma
wp_lift_head_stuck
E
Φ
e
:
...
@@ -76,61 +76,61 @@ Qed.
...
@@ -76,61 +76,61 @@ Qed.
Lemma
wp_lift_atomic_head_step_fupd
{
s
E1
E2
Φ
}
e1
:
Lemma
wp_lift_atomic_head_step_fupd
{
s
E1
E2
Φ
}
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E1
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E1
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
⌜
head_reducible
e1
σ
1
⌝
∗
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
E1
,
E2
}
▷
=
∗
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
E1
,
E2
}
▷
=
∗
state_interp
σ
2
κ
s
'
∗
state_interp
σ
2
κ
s
∗
from_option
Φ
False
(
to_val
e2
)
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
from_option
Φ
False
(
to_val
e2
)
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
s
;
E1
{{
Φ
}}.
⊢
WP
e1
@
s
;
E1
{{
Φ
}}.
Proof
.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_step_fupd
;
[
done
|].
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_step_fupd
;
[
done
|].
iIntros
(
σ
1
κ
s
)
"Hσ1"
.
iMod
(
"H"
with
"Hσ1"
)
as
"[% H]"
;
iModIntro
.
iIntros
(
σ
1
κ
κ
s
)
"Hσ1"
.
iMod
(
"H"
with
"Hσ1"
)
as
"[% H]"
;
iModIntro
.
iSplit
;
first
by
destruct
s
;
auto
.
iIntros
(
κ
κ
s'
e2
σ
2
efs
[
Hstep
->]
).
iSplit
;
first
by
destruct
s
;
auto
.
iIntros
(
e2
σ
2
efs
Hstep
).
iApply
"H"
;
eauto
.
iApply
"H"
;
eauto
.
Qed
.
Qed
.
Lemma
wp_lift_atomic_head_step
{
s
E
Φ
}
e1
:
Lemma
wp_lift_atomic_head_step
{
s
E
Φ
}
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
⌜
head_reducible
e1
σ
1
⌝
∗
▷
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
E
}=
∗
▷
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
E
}=
∗
state_interp
σ
2
κ
s
'
∗
state_interp
σ
2
κ
s
∗
from_option
Φ
False
(
to_val
e2
)
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
from_option
Φ
False
(
to_val
e2
)
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
Proof
.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_step
;
eauto
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_step
;
eauto
.
iIntros
(
σ
1
κ
s
)
"Hσ1"
.
iMod
(
"H"
with
"Hσ1"
)
as
"[% H]"
;
iModIntro
.
iIntros
(
σ
1
κ
κ
s
)
"Hσ1"
.
iMod
(
"H"
with
"Hσ1"
)
as
"[% H]"
;
iModIntro
.
iSplit
;
first
by
destruct
s
;
auto
.
iNext
.
iIntros
(
κ
κ
s'
e2
σ
2
efs
[
Hstep
->]
).
iSplit
;
first
by
destruct
s
;
auto
.
iNext
.
iIntros
(
e2
σ
2
efs
Hstep
).
iApply
"H"
;
eauto
.
iApply
"H"
;
eauto
.
Qed
.
Qed
.
Lemma
wp_lift_atomic_head_step_no_fork_fupd
{
s
E1
E2
Φ
}
e1
:
Lemma
wp_lift_atomic_
_
head_step_no_fork_fupd
{
s
E1
E2
Φ
}
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E1
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E1
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
⌜
head_reducible
e1
σ
1
⌝
∗
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
E1
,
E2
}
▷
=
∗
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
E1
,
E2
}
▷
=
∗
⌜
efs
=
[]
⌝
∗
state_interp
σ
2
κ
s
'
∗
from_option
Φ
False
(
to_val
e2
))
⌜
efs
=
[]
⌝
∗
state_interp
σ
2
κ
s
∗
from_option
Φ
False
(
to_val
e2
))
⊢
WP
e1
@
s
;
E1
{{
Φ
}}.
⊢
WP
e1
@
s
;
E1
{{
Φ
}}.
Proof
.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_head_step_fupd
;
[
done
|].
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_head_step_fupd
;
[
done
|].
iIntros
(
σ
1
κ
s
)
"Hσ1"
.
iMod
(
"H"
$!
σ
1
κ
s
with
"Hσ1"
)
as
"[$ H]"
;
iModIntro
.
iIntros
(
σ
1
κ
κ
s
)
"Hσ1"
.
iMod
(
"H"
$!
σ
1
with
"Hσ1"
)
as
"[$ H]"
;
iModIntro
.
iIntros
(
κ
κ
s'
v2
σ
2
efs
[
Hstep
->]
).
iIntros
(
v2
σ
2
efs
Hstep
).
iMod
(
"H"
$!
κ
κ
s'
v2
σ
2
efs
with
"[# //]"
)
as
"H"
.
iMod
(
"H"
$!
v2
σ
2
efs
with
"[# //]"
)
as
"H"
.
iIntros
"!> !>"
.
iMod
"H"
as
"(% & $ & $)"
;
subst
;
auto
.
iIntros
"!> !>"
.
iMod
"H"
as
"(% & $ & $)"
;
subst
;
auto
.
Qed
.
Qed
.
Lemma
wp_lift_atomic_head_step_no_fork
{
s
E
Φ
}
e1
:
Lemma
wp_lift_atomic_head_step_no_fork
{
s
E
Φ
}
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
⌜
head_reducible
e1
σ
1
⌝
∗
▷
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
E
}=
∗
▷
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
E
}=
∗
⌜
efs
=
[]
⌝
∗
state_interp
σ
2
κ
s
'
∗
from_option
Φ
False
(
to_val
e2
))
⌜
efs
=
[]
⌝
∗
state_interp
σ
2
κ
s
∗
from_option
Φ
False
(
to_val
e2
))
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
Proof
.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_head_step
;
eauto
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_head_step
;
eauto
.
iIntros
(
σ
1
κ
s
)
"Hσ1"
.
iMod
(
"H"
$!
σ
1
with
"Hσ1"
)
as
"[$ H]"
;
iModIntro
.
iIntros
(
σ
1
κ
κ
s
)
"Hσ1"
.
iMod
(
"H"
$!
σ
1
with
"Hσ1"
)
as
"[$ H]"
;
iModIntro
.
iNext
;
iIntros
(
κ
κ
s'
v2
σ
2
efs
Hstep
).
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
).
iMod
(
"H"
$!
κ
κ
s'
v2
σ
2
efs
with
"[# //]"
)
as
"(% & $ & $)"
.
subst
;
auto
.
iMod
(
"H"
$!
v2
σ
2
efs
with
"[# //]"
)
as
"(% & $ & $)"
.
subst
;
auto
.
Qed
.
Qed
.
Lemma
wp_lift_pure_det_head_step
{
s
E
E'
Φ
}
e1
e2
efs
:
Lemma
wp_lift_pure_det_head_step
{
s
E
E'
Φ
}
e1
e2
efs
:
...
...
theories/program_logic/lifting.v
View file @
4ec14182
...
@@ -15,15 +15,15 @@ Hint Resolve reducible_no_obs_reducible.
...
@@ -15,15 +15,15 @@ Hint Resolve reducible_no_obs_reducible.
Lemma
wp_lift_step_fupd
s
E
Φ
e1
:
Lemma
wp_lift_step_fupd
s
E
Φ
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E
,
∅
}=
∗
(
∀
σ
1
κ
κ
s
,
state_interp
σ
1
(
cons_obs
κ
κ
s
)
={
E
,
∅
}=
∗
⌜
if
s
is
NotStuck
then
reducible
e1
σ
1
else
True
⌝
∗
⌜
if
s
is
NotStuck
then
reducible
e1
σ
1
else
True
⌝
∗
∀
κ
κ
s'
e2
σ
2
efs
,
⌜
prim_step
e1
σ
1
κ
e2
σ
2
efs
∧
κ
s
=
cons_obs
κ
κ
s'
⌝
={
∅
,
∅
,
E
}
▷
=
∗
∀
e2
σ
2
efs
,
⌜
prim_step
e1
σ
1
κ
e2
σ
2
efs
⌝
={
∅
,
∅
,
E
}
▷
=
∗
state_interp
σ
2
κ
s
'
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
state_interp
σ
2
κ
s
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
⊢
WP
e1
@
s
;
E
{{
Φ
}}.
Proof
.
Proof
.
rewrite
wp_unfold
/
wp_pre
=>->.
iIntros
"H"
(
σ
1
κ
s
)
"Hσ"
.
rewrite
wp_unfold
/
wp_pre
=>->.
iIntros
"H"
(
σ
1
κ
κ
s
)
"Hσ"
.
iMod
(
"H"
with
"Hσ"
)
as
"(%&H)"
.
iModIntro
.
iSplit
.
by
destruct
s
.
iMod
(
"H"
with
"Hσ"
)
as
"(%&H)"
.
iModIntro
.
iSplit
.
by
destruct
s
.
iIntros
(????
?
[?
->]
).
iApply
"H"
.
eauto
.
iIntros
(????).
iApply
"H"
.
eauto
.
Qed
.
Qed
.
Lemma
wp_lift_stuck
E
Φ
e
:
Lemma
wp_lift_stuck
E
Φ
e
:
...
@@ -31,21 +31,21 @@ Lemma wp_lift_stuck E Φ e :
...
@@ -31,21 +31,21 @@ Lemma wp_lift_stuck E Φ e :
(
∀
σ
κ
s
,
state_interp
σ
κ
s
={
E
,
∅
}=
∗
⌜
stuck
e
σ⌝
)
(
∀
σ
κ
s
,
state_interp
σ
κ
s
={
E
,
∅
}=
∗
⌜
stuck
e
σ⌝
)
⊢
WP
e
@
E
?{{
Φ
}}.
⊢
WP
e
@
E
?{{
Φ
}}.
Proof
.
Proof
.
rewrite
wp_unfold
/
wp_pre
=>->.
iIntros
"H"
(
σ
1
κ
s
)
"Hσ"
.
rewrite
wp_unfold
/
wp_pre
=>->.
iIntros
"H"
(
σ
1
κ
κ
s
)
"Hσ"
.
iMod
(
"H"
with
"Hσ"
)
as
%[?
Hirr
].
iModIntro
.
iSplit
;
first
done
.
iMod
(
"H"
with
"Hσ"
)
as
%[?
Hirr
].
iModIntro
.
iSplit
;
first
done
.
iIntros
(
κ
?
e2
σ
2
efs
[?
->]
).
by
case
:
(
Hirr
κ
e2
σ
2
efs
).
iIntros
(
e2
σ
2
efs
?
).
by
case
:
(
Hirr
κ
e2
σ
2
efs
).
Qed
.
Qed
.
(** Derived lifting lemmas. *)
(** Derived lifting lemmas. *)
Lemma
wp_lift_step
s
E
Φ
e1
:
Lemma
wp_lift_step
s
E
Φ
e1
:
to_val
e1
=
None
→
to_val
e1
=
None
→
(
∀
σ
1
κ
s
,
state_interp
σ
1
κ
s
={
E
,
∅
}=
∗