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
5b4cd196
Commit
5b4cd196
authored
Dec 06, 2016
by
David Swasey
Browse files
Add progress bit to wp (and slightly generalize wp_step, wp_safe).
parent
726366bb
Changes
15
Expand all
Hide whitespace changes
Inline
Side-by-side
naming.txt
View file @
5b4cd196
...
...
@@ -14,7 +14,7 @@ l
m : iGst = ghost state
n
o
p
p
: progress bits
q
r : iRes = resources
s : state (STSs)
...
...
theories/heap_lang/adequacy.v
View file @
5b4cd196
...
...
@@ -16,7 +16,7 @@ Proof. solve_inG. Qed.
Definition
heap_adequacy
Σ
`
{
heapPreG
Σ
}
e
σ
φ
:
(
∀
`
{
heapG
Σ
},
WP
e
{{
v
,
⌜φ
v
⌝
}}%
I
)
→
adequate
e
σ
φ
.
adequate
true
e
σ
φ
.
Proof
.
intros
Hwp
;
eapply
(
wp_adequacy
_
_
)
;
iIntros
(?)
""
.
iMod
(
own_alloc
(
●
to_gen_heap
σ
))
as
(
γ
)
"Hh"
.
...
...
theories/heap_lang/lifting.v
View file @
5b4cd196
...
...
@@ -76,7 +76,7 @@ Lemma wp_fork E e Φ :
▷
Φ
(
LitV
LitUnit
)
∗
▷
WP
e
{{
_
,
True
}}
⊢
WP
Fork
e
@
E
{{
Φ
}}.
Proof
.
rewrite
-(
wp_lift_pure_det_head_step
(
Fork
e
)
(
Lit
LitUnit
)
[
e
])
//=
;
eauto
.
-
by
rewrite
-
step_fupd_intro
//
later_sep
-(
wp_value
_
_
(
Lit
_
))
//
right_id
.
-
by
rewrite
-
step_fupd_intro
//
later_sep
-(
wp_value
_
_
_
(
Lit
_
))
//
right_id
.
-
intros
;
inv_head_step
;
eauto
.
Qed
.
...
...
theories/heap_lang/proofmode.v
View file @
5b4cd196
...
...
@@ -27,7 +27,7 @@ Ltac wp_value_head := eapply tac_wp_value; [apply _|lazy beta].
Tactic
Notation
"wp_pure"
open_constr
(
efoc
)
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
reshape_expr
e
ltac
:
(
fun
K
e'
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
reshape_expr
e
ltac
:
(
fun
K
e'
=>
unify
e'
efoc
;
eapply
(
tac_wp_pure
K
)
;
[
simpl
;
apply
_
(* PureExec *)
...
...
@@ -66,7 +66,7 @@ Ltac wp_bind_core K :=
Tactic
Notation
"wp_bind"
open_constr
(
efoc
)
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
reshape_expr
e
ltac
:
(
fun
K
e'
=>
unify
e'
efoc
;
wp_bind_core
K
)
||
fail
"wp_bind: cannot find"
efoc
"in"
e
|
_
=>
fail
"wp_bind: not a 'wp'"
...
...
@@ -151,7 +151,7 @@ End heap.
Tactic
Notation
"wp_apply"
open_constr
(
lem
)
:
=
iPoseProofCore
lem
as
false
true
(
fun
H
=>
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
reshape_expr
e
ltac
:
(
fun
K
e'
=>
wp_bind_core
K
;
iApplyHyp
H
;
try
iNext
;
simpl
)
||
lazymatch
iTypeOf
H
with
...
...
@@ -163,7 +163,7 @@ Tactic Notation "wp_apply" open_constr(lem) :=
Tactic
Notation
"wp_alloc"
ident
(
l
)
"as"
constr
(
H
)
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
first
[
reshape_expr
e
ltac
:
(
fun
K
e'
=>
eapply
(
tac_wp_alloc
_
_
_
H
K
)
;
[
apply
_
|..])
...
...
@@ -182,7 +182,7 @@ Tactic Notation "wp_alloc" ident(l) :=
Tactic
Notation
"wp_load"
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
first
[
reshape_expr
e
ltac
:
(
fun
K
e'
=>
eapply
(
tac_wp_load
_
_
_
_
K
))
|
fail
1
"wp_load: cannot find 'Load' in"
e
]
;
...
...
@@ -196,7 +196,7 @@ Tactic Notation "wp_load" :=
Tactic
Notation
"wp_store"
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
first
[
reshape_expr
e
ltac
:
(
fun
K
e'
=>
eapply
(
tac_wp_store
_
_
_
_
_
K
)
;
[
apply
_
|..])
...
...
@@ -212,7 +212,7 @@ Tactic Notation "wp_store" :=
Tactic
Notation
"wp_cas_fail"
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
first
[
reshape_expr
e
ltac
:
(
fun
K
e'
=>
eapply
(
tac_wp_cas_fail
_
_
_
_
K
)
;
[
apply
_
|
apply
_
|..])
...
...
@@ -228,7 +228,7 @@ Tactic Notation "wp_cas_fail" :=
Tactic
Notation
"wp_cas_suc"
:
=
iStartProof
;
lazymatch
goal
with
|
|-
envs_entails
_
(
wp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
wp
true
?E
?e
?Q
)
=>
first
[
reshape_expr
e
ltac
:
(
fun
K
e'
=>
eapply
(
tac_wp_cas_suc
_
_
_
_
_
K
)
;
[
apply
_
|
apply
_
|..])
...
...
theories/heap_lang/tactics.v
View file @
5b4cd196
...
...
@@ -187,11 +187,10 @@ Definition is_atomic (e : expr) :=
|
App
(
Rec
_
_
(
Lit
_
))
(
Lit
_
)
=>
true
|
_
=>
false
end
.
Lemma
is_atomic_correct
e
:
is_atomic
e
→
Atomic
(
to_expr
e
).
Lemma
is_atomic_correct
e
:
is_atomic
e
→
Strongly
Atomic
(
to_expr
e
).
Proof
.
intros
He
.
apply
ectx_language_atomic
.
-
intros
σ
e'
σ
'
ef
Hstep
;
simpl
in
*.
apply
language
.
val_irreducible
;
revert
Hstep
.
intros
He
.
apply
ectx_language_strongly_atomic
.
-
intros
σ
e'
σ
'
ef
Hstep
;
simpl
in
*.
revert
Hstep
.
destruct
e
=>
//=
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
inversion
1
;
simplify_eq
/=
;
rewrite
?to_of_val
;
eauto
.
unfold
subst'
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
eauto
.
...
...
@@ -227,11 +226,11 @@ Hint Extern 10 (AsVal _) => solve_as_val : typeclass_instances.
Ltac
solve_atomic
:
=
match
goal
with
|
|-
Atomic
?e
=>
let
e'
:
=
W
.
of_expr
e
in
change
(
Atomic
(
W
.
to_expr
e'
))
;
|
|-
Strongly
Atomic
?e
=>
let
e'
:
=
W
.
of_expr
e
in
change
(
Strongly
Atomic
(
W
.
to_expr
e'
))
;
apply
W
.
is_atomic_correct
;
vm_compute
;
exact
I
end
.
Hint
Extern
10
(
Atomic
_
)
=>
solve_atomic
:
typeclass_instances
.
Hint
Extern
10
(
Strongly
Atomic
_
)
=>
solve_atomic
:
typeclass_instances
.
(** Substitution *)
Ltac
simpl_subst
:
=
...
...
theories/program_logic/adequacy.v
View file @
5b4cd196
...
...
@@ -34,23 +34,24 @@ Proof.
Qed
.
(* Program logic adequacy *)
Record
adequate
{
Λ
}
(
e1
:
expr
Λ
)
(
σ
1
:
state
Λ
)
(
φ
:
val
Λ
→
Prop
)
:
=
{
Record
adequate
{
Λ
}
(
p
:
bool
)
(
e1
:
expr
Λ
)
(
σ
1
:
state
Λ
)
(
φ
:
val
Λ
→
Prop
)
:
=
{
adequate_result
t2
σ
2
v2
:
rtc
step
([
e1
],
σ
1
)
(
of_val
v2
::
t2
,
σ
2
)
→
φ
v2
;
adequate_safe
t2
σ
2 e2
:
p
→
rtc
step
([
e1
],
σ
1
)
(
t2
,
σ
2
)
→
e2
∈
t2
→
(
is_Some
(
to_val
e2
)
∨
reducibl
e
e2
σ
2
)
e2
∈
t2
→
progressiv
e
e2
σ
2
}.
Theorem
adequate_tp_safe
{
Λ
}
(
e1
:
expr
Λ
)
t2
σ
1
σ
2
φ
:
adequate
e1
σ
1
φ
→
adequate
true
e1
σ
1
φ
→
rtc
step
([
e1
],
σ
1
)
(
t2
,
σ
2
)
→
Forall
(
λ
e
,
is_Some
(
to_val
e
))
t2
∨
∃
t3
σ
3
,
step
(
t2
,
σ
2
)
(
t3
,
σ
3
).
Proof
.
intros
Had
?.
destruct
(
decide
(
Forall
(
λ
e
,
is_Some
(
to_val
e
))
t2
))
as
[|
Ht2
]
;
[
by
left
|].
apply
(
not_Forall_Exists
_
),
Exists_exists
in
Ht2
;
destruct
Ht2
as
(
e2
&?&
He2
).
destruct
(
adequate_safe
e1
σ
1
φ
Had
t2
σ
2 e2
)
as
[?|(
e3
&
σ
3
&
efs
&?)]
;
destruct
(
adequate_safe
true
e1
σ
1
φ
Had
t2
σ
2 e2
)
as
[?|(
e3
&
σ
3
&
efs
&?)]
;
rewrite
?eq_None_not_Some
;
auto
.
{
exfalso
.
eauto
.
}
destruct
(
elem_of_list_split
t2
e2
)
as
(
t2'
&
t2''
&->)
;
auto
.
...
...
@@ -58,19 +59,22 @@ Proof.
Qed
.
Section
adequacy
.
Context
`
{
irisG
Λ
Σ
}.
Context
`
{
irisG
Λ
Σ
}
(
p
:
bool
)
.
Implicit
Types
e
:
expr
Λ
.
Implicit
Types
P
Q
:
iProp
Σ
.
Implicit
Types
Φ
:
val
Λ
→
iProp
Σ
.
Implicit
Types
Φ
s
:
list
(
val
Λ
→
iProp
Σ
).
Notation
world
σ
:
=
(
wsat
∗
ownE
⊤
∗
state_interp
σ
)%
I
.
Notation
world'
E
σ
:
=
(
wsat
∗
ownE
E
∗
state_interp
σ
)%
I
.
Notation
world
σ
:
=
(
world'
⊤
σ
).
Notation
wptp
t
:
=
([
∗
list
]
ef
∈
t
,
WP
ef
{{
_
,
True
}})%
I
.
Notation
wp'
E
e
Φ
:
=
(
WP
e
@
p
;
E
{{
Φ
}})%
I
.
Notation
wp
e
Φ
:
=
(
wp'
⊤
e
Φ
).
Notation
wptp
t
:
=
([
∗
list
]
ef
∈
t
,
WP
ef
@
p
;
⊤
{{
_
,
True
}})%
I
.
Lemma
wp_step
e1
σ
1 e2
σ
2
efs
Φ
:
Lemma
wp_step
E
e1
σ
1 e2
σ
2
efs
Φ
:
prim_step
e1
σ
1 e2
σ
2
efs
→
world
σ
1
∗
WP
e1
{{
Φ
}}
==
∗
▷
|==>
◇
(
world
σ
2
∗
WP
e2
{{
Φ
}}
∗
wptp
efs
).
world
'
E
σ
1
∗
wp'
E
e1
Φ
==
∗
▷
|==>
◇
(
world
'
E
σ
2
∗
wp'
E
e2
Φ
∗
wptp
efs
).
Proof
.
rewrite
{
1
}
wp_unfold
/
wp_pre
.
iIntros
(?)
"[(Hw & HE & Hσ) H]"
.
rewrite
(
val_stuck
e1
σ
1 e2
σ
2
efs
)
//
fupd_eq
/
fupd_def
.
...
...
@@ -81,8 +85,8 @@ Qed.
Lemma
wptp_step
e1
t1
t2
σ
1
σ
2
Φ
:
step
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
world
σ
1
∗
WP
e1
{{
Φ
}}
∗
wptp
t1
==
∗
∃
e2
t2'
,
⌜
t2
=
e2
::
t2'
⌝
∗
▷
|==>
◇
(
world
σ
2
∗
WP
e2
{{
Φ
}}
∗
wptp
t2'
).
world
σ
1
∗
wp
e1
Φ
∗
wptp
t1
==
∗
∃
e2
t2'
,
⌜
t2
=
e2
::
t2'
⌝
∗
▷
|==>
◇
(
world
σ
2
∗
wp
e2
Φ
∗
wptp
t2'
).
Proof
.
iIntros
(
Hstep
)
"(HW & He & Ht)"
.
destruct
Hstep
as
[
e1'
σ
1
'
e2'
σ
2
'
efs
[|?
t1'
]
t2'
??
Hstep
]
;
simplify_eq
/=.
...
...
@@ -95,9 +99,9 @@ Qed.
Lemma
wptp_steps
n
e1
t1
t2
σ
1
σ
2
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
world
σ
1
∗
WP
e1
{{
Φ
}}
∗
wptp
t1
⊢
world
σ
1
∗
wp
e1
Φ
∗
wptp
t1
⊢
Nat
.
iter
(
S
n
)
(
λ
P
,
|==>
▷
P
)
(
∃
e2
t2'
,
⌜
t2
=
e2
::
t2'
⌝
∗
world
σ
2
∗
WP
e2
{{
Φ
}}
∗
wptp
t2'
).
⌜
t2
=
e2
::
t2'
⌝
∗
world
σ
2
∗
wp
e2
Φ
∗
wptp
t2'
).
Proof
.
revert
e1
t1
t2
σ
1
σ
2
;
simpl
;
induction
n
as
[|
n
IH
]=>
e1
t1
t2
σ
1
σ
2
/=.
{
inversion_clear
1
;
iIntros
"?"
;
eauto
10
.
}
...
...
@@ -121,7 +125,7 @@ Qed.
Lemma
wptp_result
n
e1
t1
v2
t2
σ
1
σ
2
φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
of_val
v2
::
t2
,
σ
2
)
→
world
σ
1
∗
WP
e1
{{
v
,
⌜φ
v
⌝
}}
∗
wptp
t1
⊢
▷
^(
S
(
S
n
))
⌜φ
v2
⌝
.
world
σ
1
∗
wp
e1
(
λ
v
,
⌜φ
v
⌝
)
∗
wptp
t1
⊢
▷
^(
S
(
S
n
))
⌜φ
v2
⌝
.
Proof
.
intros
.
rewrite
wptp_steps
//
laterN_later
.
apply
:
bupd_iter_laterN_mono
.
iDestruct
1
as
(
e2
t2'
?)
"((Hw & HE & _) & H & _)"
;
simplify_eq
.
...
...
@@ -129,18 +133,20 @@ Proof.
iMod
(
"H"
with
"[Hw HE]"
)
as
">(_ & _ & $)"
;
iFrame
;
auto
.
Qed
.
Lemma
wp_safe
e
σ
Φ
:
world
σ
-
∗
WP
e
{{
Φ
}}
==
∗
▷
⌜
is_Some
(
to_val
e
)
∨
reducible
e
σ
⌝
.
Lemma
wp_safe
E
e
σ
Φ
:
world
'
E
σ
-
∗
wp'
E
e
Φ
==
∗
▷
⌜
if
p
then
progressive
e
σ
else
True
⌝
.
Proof
.
rewrite
wp_unfold
/
wp_pre
.
iIntros
"(Hw&HE&Hσ) H"
.
destruct
(
to_val
e
)
as
[
v
|]
eqn
:
?
;
[
eauto
10
|].
rewrite
fupd_eq
.
iMod
(
"H"
with
"Hσ [-]"
)
as
">(?&?&%&?)"
;
eauto
10
with
iFrame
.
destruct
(
to_val
e
)
as
[
v
|]
eqn
:
?.
{
destruct
p
;
last
done
.
iIntros
"!> !> !%"
.
left
.
by
exists
v
.
}
rewrite
fupd_eq
.
iMod
(
"H"
with
"Hσ [-]"
)
as
">(?&?&%&?)"
;
first
by
iFrame
.
destruct
p
;
last
done
.
iIntros
"!> !> !%"
.
by
right
.
Qed
.
Lemma
wptp_safe
n
e1
e2
t1
t2
σ
1
σ
2
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
e2
∈
t2
→
world
σ
1
∗
WP
e1
{{
Φ
}}
∗
wptp
t1
⊢
▷
^(
S
(
S
n
))
⌜
i
s_Some
(
to_val
e2
)
∨
reducible
e2
σ
2
⌝
.
world
σ
1
∗
wp
e1
Φ
∗
wptp
t1
⊢
▷
^(
S
(
S
n
))
⌜
i
f
p
then
progressive
e2
σ
2
else
True
⌝
.
Proof
.
intros
?
He2
.
rewrite
wptp_steps
//
laterN_later
.
apply
:
bupd_iter_laterN_mono
.
iDestruct
1
as
(
e2'
t2'
?)
"(Hw & H & Htp)"
;
simplify_eq
.
...
...
@@ -151,7 +157,7 @@ Qed.
Lemma
wptp_invariance
n
e1
e2
t1
t2
σ
1
σ
2
φ
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
(
state_interp
σ
2
={
⊤
,
∅
}=
∗
⌜φ⌝
)
∗
world
σ
1
∗
WP
e1
{{
Φ
}}
∗
wptp
t1
(
state_interp
σ
2
={
⊤
,
∅
}=
∗
⌜φ⌝
)
∗
world
σ
1
∗
wp
e1
Φ
∗
wptp
t1
⊢
▷
^(
S
(
S
n
))
⌜φ⌝
.
Proof
.
intros
?.
rewrite
wptp_steps
//
bupd_iter_frame_l
laterN_later
.
...
...
@@ -162,12 +168,12 @@ Proof.
Qed
.
End
adequacy
.
Theorem
wp_adequacy
Σ
Λ
`
{
invPreG
Σ
}
e
σ
φ
:
Theorem
wp_adequacy
Σ
Λ
`
{
invPreG
Σ
}
p
e
σ
φ
:
(
∀
`
{
Hinv
:
invG
Σ
},
(|={
⊤
}=>
∃
stateI
:
state
Λ
→
iProp
Σ
,
let
_
:
irisG
Λ
Σ
:
=
IrisG
_
_
Hinv
stateI
in
stateI
σ
∗
WP
e
{{
v
,
⌜φ
v
⌝
}})%
I
)
→
adequate
e
σ
φ
.
stateI
σ
∗
WP
e
@
p
;
⊤
{{
v
,
⌜φ
v
⌝
}})%
I
)
→
adequate
p
e
σ
φ
.
Proof
.
intros
Hwp
;
split
.
-
intros
t2
σ
2
v2
[
n
?]%
rtc_nsteps
.
...
...
@@ -176,19 +182,19 @@ Proof.
rewrite
fupd_eq
in
Hwp
;
iMod
(
Hwp
with
"[$Hw $HE]"
)
as
">(Hw & HE & Hwp)"
.
iDestruct
"Hwp"
as
(
Istate
)
"[HI Hwp]"
.
iApply
(@
wptp_result
_
_
(
IrisG
_
_
Hinv
Istate
))
;
eauto
with
iFrame
.
-
intros
t2
σ
2 e2
[
n
?]%
rtc_nsteps
?.
-
destruct
p
;
last
done
.
intros
t2
σ
2 e2
_
[
n
?]%
rtc_nsteps
?.
eapply
(
soundness
(
M
:
=
iResUR
Σ
)
_
(
S
(
S
n
))).
iMod
wsat_alloc
as
(
Hinv
)
"[Hw HE]"
.
rewrite
fupd_eq
in
Hwp
;
iMod
(
Hwp
with
"[$Hw $HE]"
)
as
">(Hw & HE & Hwp)"
.
iDestruct
"Hwp"
as
(
Istate
)
"[HI Hwp]"
.
iApply
(@
wptp_safe
_
_
(
IrisG
_
_
Hinv
Istate
))
;
eauto
with
iFrame
.
iApply
(@
wptp_safe
_
_
(
IrisG
_
_
Hinv
Istate
)
true
)
;
eauto
with
iFrame
.
Qed
.
Theorem
wp_invariance
Σ
Λ
`
{
invPreG
Σ
}
e
σ
1
t2
σ
2
φ
:
Theorem
wp_invariance
Σ
Λ
`
{
invPreG
Σ
}
p
e
σ
1
t2
σ
2
φ
:
(
∀
`
{
Hinv
:
invG
Σ
},
(|={
⊤
}=>
∃
stateI
:
state
Λ
→
iProp
Σ
,
let
_
:
irisG
Λ
Σ
:
=
IrisG
_
_
Hinv
stateI
in
stateI
σ
1
∗
WP
e
{{
_
,
True
}}
∗
(
stateI
σ
2
={
⊤
,
∅
}=
∗
⌜φ⌝
))%
I
)
→
stateI
σ
1
∗
WP
e
@
p
;
⊤
{{
_
,
True
}}
∗
(
stateI
σ
2
={
⊤
,
∅
}=
∗
⌜φ⌝
))%
I
)
→
rtc
step
([
e
],
σ
1
)
(
t2
,
σ
2
)
→
φ
.
Proof
.
...
...
theories/program_logic/ectx_language.v
View file @
5b4cd196
...
...
@@ -127,6 +127,16 @@ Section ectx_language.
rewrite
fill_empty
.
eapply
Hatomic_step
.
by
rewrite
fill_empty
.
Qed
.
Lemma
ectx_language_strongly_atomic
e
:
(
∀
σ
e'
σ
'
efs
,
head_step
e
σ
e'
σ
'
efs
→
is_Some
(
to_val
e'
))
→
sub_values
e
→
StronglyAtomic
e
.
Proof
.
intros
Hatomic_step
Hatomic_fill
σ
e'
σ
'
efs
[
K
e1'
e2'
->
->
Hstep
].
assert
(
K
=
empty_ectx
)
as
->
by
eauto
10
using
val_stuck
.
rewrite
fill_empty
.
eapply
Hatomic_step
.
by
rewrite
fill_empty
.
Qed
.
Lemma
head_reducible_prim_step
e1
σ
1 e2
σ
2
efs
:
head_reducible
e1
σ
1
→
prim_step
e1
σ
1 e2
σ
2
efs
→
...
...
theories/program_logic/ectx_lifting.v
View file @
5b4cd196
...
...
@@ -12,61 +12,93 @@ Implicit Types v : val.
Implicit
Types
e
:
expr
.
Hint
Resolve
head_prim_reducible
head_reducible_prim_step
.
Lemma
wp_ectx_bind
{
E
Φ
}
K
e
:
WP
e
@
E
{{
v
,
WP
fill
K
(
of_val
v
)
@
E
{{
Φ
}}
}}
⊢
WP
fill
K
e
@
E
{{
Φ
}}.
Definition
head_progressive
(
e
:
expr
)
(
σ
:
state
)
:
=
is_Some
(
to_val
e
)
∨
∃
K
e'
,
e
=
fill
K
e'
∧
head_reducible
e'
σ
.
Lemma
progressive_head_progressive
e
σ
:
progressive
e
σ
→
head_progressive
e
σ
.
Proof
.
case
=>[?|
Hred
]
;
first
by
left
.
right
.
move
:
Hred
=>
[]
e'
[]
σ
'
[]
efs
[]
K
e1'
e2'
EQ
EQ'
Hstep
.
subst
.
exists
K
,
e1'
.
split
;
first
done
.
by
exists
e2'
,
σ
'
,
efs
.
Qed
.
Hint
Resolve
progressive_head_progressive
.
Lemma
wp_ectx_bind
{
p
E
e
}
K
Φ
:
WP
e
@
p
;
E
{{
v
,
WP
fill
K
(
of_val
v
)
@
p
;
E
{{
Φ
}}
}}
⊢
WP
fill
K
e
@
p
;
E
{{
Φ
}}.
Proof
.
apply
:
weakestpre
.
wp_bind
.
Qed
.
Lemma
wp_ectx_bind_inv
{
E
Φ
}
K
e
:
WP
fill
K
e
@
E
{{
Φ
}}
⊢
WP
e
@
E
{{
v
,
WP
fill
K
(
of_val
v
)
@
E
{{
Φ
}}
}}.
Lemma
wp_ectx_bind_inv
{
p
E
Φ
}
K
e
:
WP
fill
K
e
@
p
;
E
{{
Φ
}}
⊢
WP
e
@
p
;
E
{{
v
,
WP
fill
K
(
of_val
v
)
@
p
;
E
{{
Φ
}}
}}.
Proof
.
apply
:
weakestpre
.
wp_bind_inv
.
Qed
.
Lemma
wp_lift_head_step
{
E
Φ
}
e1
:
Lemma
wp_lift_head_step
{
p
E
Φ
}
e1
:
to_val
e1
=
None
→
(
∀
σ
1
,
state_interp
σ
1
={
E
,
∅
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
▷
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1 e2
σ
2
efs
⌝
={
∅
,
E
}=
∗
state_interp
σ
2
∗
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
state_interp
σ
2
∗
WP
e2
@
p
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
p
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
.
iIntros
(?)
"H"
.
iApply
(
wp_lift_step
E
)
=>//.
iIntros
(
σ
1
)
"Hσ"
.
iIntros
(?)
"H"
.
iApply
wp_lift_step
=>//.
iIntros
(
σ
1
)
"Hσ"
.
iMod
(
"H"
$!
σ
1
with
"Hσ"
)
as
"[% H]"
;
iModIntro
.
iSplit
;
first
by
eauto
.
iNext
.
iIntros
(
e2
σ
2
efs
)
"%"
.
iApply
"H"
.
by
eauto
.
Qed
.
Lemma
wp_lift_pure_head_step
{
E
E'
Φ
}
e1
:
Lemma
wp_lift_head_stuck
E
Φ
e
:
to_val
e
=
None
→
(
∀
σ
,
state_interp
σ
={
E
,
∅
}=
∗
⌜
¬
head_progressive
e
σ⌝
)
⊢
WP
e
@
E
?{{
Φ
}}.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_stuck
;
first
done
.
iIntros
(
σ
)
"Hσ"
.
iMod
(
"H"
$!
_
with
"Hσ"
)
as
"%"
.
iModIntro
.
by
auto
.
Qed
.
Lemma
wp_lift_pure_head_step
{
p
E
E'
Φ
}
e1
:
(
∀
σ
1
,
head_reducible
e1
σ
1
)
→
(
∀
σ
1 e2
σ
2
efs
,
head_step
e1
σ
1 e2
σ
2
efs
→
σ
1
=
σ
2
)
→
(|={
E
,
E'
}
▷
=>
∀
e2
efs
σ
,
⌜
head_step
e1
σ
e2
σ
efs
⌝
→
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
WP
e2
@
p
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
p
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
using
Hinh
.
iIntros
(??)
"H"
.
iApply
wp_lift_pure_step
;
eauto
.
iApply
(
step_fupd_wand
with
"H"
)
;
iIntros
"H"
.
iIntros
(????).
iApply
"H"
;
eauto
.
Qed
.
Lemma
wp_lift_atomic_head_step
{
E
Φ
}
e1
:
Lemma
wp_lift_pure_head_stuck
`
{
Inhabited
state
}
E
Φ
e
:
to_val
e
=
None
→
(
∀
K
e1
σ
1 e2
σ
2
efs
,
e
=
fill
K
e1
→
¬
head_step
e1
σ
1 e2
σ
2
efs
)
→
WP
e
@
E
?{{
Φ
}}%
I
.
Proof
.
iIntros
(
Hnv
Hnstep
).
iApply
wp_lift_head_stuck
;
first
done
.
iIntros
(
σ
)
"_"
.
iMod
(
fupd_intro_mask'
E
∅
)
as
"_"
;
first
set_solver
.
iModIntro
.
iPureIntro
.
case
;
first
by
rewrite
Hnv
;
case
.
move
=>[]
K
[]
e1
[]
Hfill
[]
e2
[]
σ
2
[]
efs
/=
Hstep
.
exact
:
Hnstep
.
Qed
.
Lemma
wp_lift_atomic_head_step
{
p
E
Φ
}
e1
:
to_val
e1
=
None
→
(
∀
σ
1
,
state_interp
σ
1
={
E
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
▷
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1 e2
σ
2
efs
⌝
={
E
}=
∗
state_interp
σ
2
∗
default
False
(
to_val
e2
)
Φ
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
default
False
(
to_val
e2
)
Φ
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
p
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_step
;
eauto
.
iIntros
(
σ
1
)
"Hσ1"
.
iMod
(
"H"
$!
σ
1
with
"Hσ1"
)
as
"[% H]"
;
iModIntro
.
iSplit
;
first
by
eauto
.
iNext
.
iIntros
(
e2
σ
2
efs
)
"%"
.
iApply
"H"
;
auto
.
Qed
.
Lemma
wp_lift_atomic_head_step_no_fork
{
E
Φ
}
e1
:
Lemma
wp_lift_atomic_head_step_no_fork
{
p
E
Φ
}
e1
:
to_val
e1
=
None
→
(
∀
σ
1
,
state_interp
σ
1
={
E
}=
∗
⌜
head_reducible
e1
σ
1
⌝
∗
▷
∀
e2
σ
2
efs
,
⌜
head_step
e1
σ
1 e2
σ
2
efs
⌝
={
E
}=
∗
⌜
efs
=
[]
⌝
∗
state_interp
σ
2
∗
default
False
(
to_val
e2
)
Φ
)
⊢
WP
e1
@
E
{{
Φ
}}.
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
.
iIntros
(?)
"H"
.
iApply
wp_lift_atomic_head_step
;
eauto
.
iIntros
(
σ
1
)
"Hσ1"
.
iMod
(
"H"
$!
σ
1
with
"Hσ1"
)
as
"[$ H]"
;
iModIntro
.
...
...
@@ -74,32 +106,32 @@ Proof.
iMod
(
"H"
$!
v2
σ
2
efs
with
"[# //]"
)
as
"(% & $ & $)"
;
subst
;
auto
.
Qed
.
Lemma
wp_lift_pure_det_head_step
{
E
E'
Φ
}
e1
e2
efs
:
Lemma
wp_lift_pure_det_head_step
{
p
E
E'
Φ
}
e1
e2
efs
:
(
∀
σ
1
,
head_reducible
e1
σ
1
)
→
(
∀
σ
1 e2
'
σ
2
efs'
,
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
efs
=
efs'
)
→
(|={
E
,
E'
}
▷
=>
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
(|={
E
,
E'
}
▷
=>
WP
e2
@
p
;
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
@
p
;
⊤
{{
_
,
True
}})
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
using
Hinh
.
eauto
using
wp_lift_pure_det_step
.
Qed
.
Lemma
wp_lift_pure_det_head_step_no_fork
{
E
E'
Φ
}
e1
e2
:
Lemma
wp_lift_pure_det_head_step_no_fork
{
p
E
E'
Φ
}
e1
e2
:
to_val
e1
=
None
→
(
∀
σ
1
,
head_reducible
e1
σ
1
)
→
(
∀
σ
1 e2
'
σ
2
efs'
,
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
[]
=
efs'
)
→
(|={
E
,
E'
}
▷
=>
WP
e2
@
E
{{
Φ
}})
⊢
WP
e1
@
E
{{
Φ
}}.
(|={
E
,
E'
}
▷
=>
WP
e2
@
p
;
E
{{
Φ
}})
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
using
Hinh
.
intros
.
rewrite
-(
wp_lift_pure_det_step
e1
e2
[])
/=
?right_id
;
eauto
.
Qed
.
Lemma
wp_lift_pure_det_head_step_no_fork'
{
E
Φ
}
e1
e2
:
Lemma
wp_lift_pure_det_head_step_no_fork'
{
p
E
Φ
}
e1
e2
:
to_val
e1
=
None
→
(
∀
σ
1
,
head_reducible
e1
σ
1
)
→
(
∀
σ
1 e2
'
σ
2
efs'
,
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
[]
=
efs'
)
→
▷
WP
e2
@
E
{{
Φ
}}
⊢
WP
e1
@
E
{{
Φ
}}.
▷
WP
e2
@
p
;
E
{{
Φ
}}
⊢
WP
e1
@
p
;
E
{{
Φ
}}.
Proof
using
Hinh
.
intros
.
rewrite
-[(
WP
e1
@
_
{{
_
}})%
I
]
wp_lift_pure_det_head_step_no_fork
//.
intros
.
rewrite
-[(
WP
e1
@
_;
_
{{
_
}})%
I
]
wp_lift_pure_det_head_step_no_fork
//.
rewrite
-
step_fupd_intro
//.
Qed
.
End
wp
.
theories/program_logic/hoare.v
View file @
5b4cd196
...
...
@@ -3,24 +3,42 @@ From iris.base_logic.lib Require Export viewshifts.
From
iris
.
proofmode
Require
Import
tactics
.
Set
Default
Proof
Using
"Type"
.
Definition
ht
`
{
irisG
Λ
Σ
}
(
E
:
coPset
)
(
P
:
iProp
Σ
)
Definition
ht
`
{
irisG
Λ
Σ
}
(
p
:
bool
)
(
E
:
coPset
)
(
P
:
iProp
Σ
)
(
e
:
expr
Λ
)
(
Φ
:
val
Λ
→
iProp
Σ
)
:
iProp
Σ
:
=
(
□
(
P
-
∗
WP
e
@
E
{{
Φ
}}))%
I
.
Instance
:
Params
(@
ht
)
4
.
(
□
(
P
-
∗
WP
e
@
p
;
E
{{
Φ
}}))%
I
.
Instance
:
Params
(@
ht
)
5
.
Notation
"{{ P } } e @ E {{ Φ } }"
:
=
(
ht
E
P
%
I
e
%
E
Φ
%
I
)
Notation
"{{ P } } e @ p ; E {{ Φ } }"
:
=
(
ht
p
E
P
%
I
e
%
E
Φ
%
I
)
(
at
level
20
,
P
,
e
,
Φ
at
level
200
,
format
"{{ P } } e @ p ; E {{ Φ } }"
)
:
C_scope
.
Notation
"{{ P } } e @ E {{ Φ } }"
:
=
(
ht
true
E
P
%
I
e
%
E
Φ
%
I
)
(
at
level
20
,
P
,
e
,
Φ
at
level
200
,
format
"{{ P } } e @ E {{ Φ } }"
)
:
C_scope
.
Notation
"{{ P } } e {{ Φ } }"
:
=
(
ht
⊤
P
%
I
e
%
E
Φ
%
I
)
Notation
"{{ P } } e @ E ? {{ Φ } }"
:
=
(
ht
false
E
P
%
I
e
%
E
Φ
%
I
)
(
at
level
20
,
P
,
e
,
Φ
at
level
200
,
format
"{{ P } } e @ E ? {{ Φ } }"
)
:
C_scope
.
Notation
"{{ P } } e {{ Φ } }"
:
=
(
ht
true
⊤
P
%
I
e
%
E
Φ
%
I
)
(
at
level
20
,
P
,
e
,
Φ
at
level
200
,
format
"{{ P } } e {{ Φ } }"
)
:
C_scope
.
Notation
"{{ P } } e ? {{ Φ } }"
:
=
(
ht
false
⊤
P
%
I
e
%
E
Φ
%
I
)
(
at
level
20
,
P
,
e
,
Φ
at
level
200
,
format
"{{ P } } e ? {{ Φ } }"
)
:
C_scope
.
Notation
"{{ P } } e @ E {{ v , Q } }"
:
=
(
ht
E
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
Notation
"{{ P } } e @ p ; E {{ v , Q } }"
:
=
(
ht
p
E
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
(
at
level
20
,
P
,
e
,
Q
at
level
200
,
format
"{{ P } } e @ p ; E {{ v , Q } }"
)
:
C_scope
.
Notation
"{{ P } } e @ E {{ v , Q } }"
:
=
(
ht
true
E
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
(
at
level
20
,
P
,
e
,
Q
at
level
200
,
format
"{{ P } } e @ E {{ v , Q } }"
)
:
C_scope
.
Notation
"{{ P } } e {{ v , Q } }"
:
=
(
ht
⊤
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
Notation
"{{ P } } e @ E ? {{ v , Q } }"
:
=
(
ht
false
E
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
(
at
level
20
,
P
,
e
,
Q
at
level
200
,
format
"{{ P } } e @ E ? {{ v , Q } }"
)
:
C_scope
.
Notation
"{{ P } } e {{ v , Q } }"
:
=
(
ht
true
⊤
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
(
at
level
20
,
P
,
e
,
Q
at
level
200
,
format
"{{ P } } e {{ v , Q } }"
)
:
C_scope
.
Notation
"{{ P } } e ? {{ v , Q } }"
:
=
(
ht
false
⊤
P
%
I
e
%
E
(
λ
v
,
Q
)%
I
)
(
at
level
20
,
P
,
e
,
Q
at
level
200
,
format
"{{ P } } e ? {{ v , Q } }"
)
:
C_scope
.