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
Jonas Kastberg
iris
Commits
936eeb48
Commit
936eeb48
authored
Jul 10, 2018
by
Marianna Rapoport
Committed by
Ralf Jung
Oct 05, 2018
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Adding initial support for observations to the definition of wp for later prophecy support
parent
53f179de
Changes
15
Show whitespace changes
Inline
Side-by-side
Showing
15 changed files
with
331 additions
and
289 deletions
+331
-289
theories/heap_lang/lang.v
theories/heap_lang/lang.v
+24
-23
theories/heap_lang/lifting.v
theories/heap_lang/lifting.v
+18
-18
theories/heap_lang/total_adequacy.v
theories/heap_lang/total_adequacy.v
+1
-1
theories/program_logic/adequacy.v
theories/program_logic/adequacy.v
+30
-28
theories/program_logic/ectx_language.v
theories/program_logic/ectx_language.v
+43
-42
theories/program_logic/ectx_lifting.v
theories/program_logic/ectx_lifting.v
+24
-24
theories/program_logic/ectxi_language.v
theories/program_logic/ectxi_language.v
+15
-14
theories/program_logic/language.v
theories/program_logic/language.v
+67
-31
theories/program_logic/lifting.v
theories/program_logic/lifting.v
+14
-14
theories/program_logic/ownp.v
theories/program_logic/ownp.v
+31
-30
theories/program_logic/total_adequacy.v
theories/program_logic/total_adequacy.v
+10
-10
theories/program_logic/total_ectx_lifting.v
theories/program_logic/total_ectx_lifting.v
+14
-14
theories/program_logic/total_lifting.v
theories/program_logic/total_lifting.v
+11
-11
theories/program_logic/total_weakestpre.v
theories/program_logic/total_weakestpre.v
+14
-14
theories/program_logic/weakestpre.v
theories/program_logic/weakestpre.v
+15
-15
No files found.
theories/heap_lang/lang.v
View file @
936eeb48
...
@@ -108,6 +108,8 @@ Inductive val :=
...
@@ -108,6 +108,8 @@ Inductive val :=
Bind
Scope
val_scope
with
val
.
Bind
Scope
val_scope
with
val
.
Inductive
observation
:
=
prophecy_observation_todo
.
Fixpoint
of_val
(
v
:
val
)
:
expr
:
=
Fixpoint
of_val
(
v
:
val
)
:
expr
:
=
match
v
with
match
v
with
|
RecV
f
x
e
=>
Rec
f
x
e
|
RecV
f
x
e
=>
Rec
f
x
e
...
@@ -415,62 +417,61 @@ Definition vals_cas_compare_safe (vl v1 : val) : Prop :=
...
@@ -415,62 +417,61 @@ Definition vals_cas_compare_safe (vl v1 : val) : Prop :=
val_is_unboxed
vl
∨
val_is_unboxed
v1
.
val_is_unboxed
vl
∨
val_is_unboxed
v1
.
Arguments
vals_cas_compare_safe
!
_
!
_
/.
Arguments
vals_cas_compare_safe
!
_
!
_
/.
Inductive
head_step
:
expr
→
state
→
expr
→
state
→
list
(
expr
)
→
Prop
:
=
Inductive
head_step
:
expr
→
state
→
option
observation
->
expr
→
state
→
list
(
expr
)
→
Prop
:
=
|
BetaS
f
x
e1
e2
v2
e'
σ
:
|
BetaS
f
x
e1
e2
v2
e'
σ
:
to_val
e2
=
Some
v2
→
to_val
e2
=
Some
v2
→
Closed
(
f
:
b
:
x
:
b
:
[])
e1
→
Closed
(
f
:
b
:
x
:
b
:
[])
e1
→
e'
=
subst'
x
(
of_val
v2
)
(
subst'
f
(
Rec
f
x
e1
)
e1
)
→
e'
=
subst'
x
(
of_val
v2
)
(
subst'
f
(
Rec
f
x
e1
)
e1
)
→
head_step
(
App
(
Rec
f
x
e1
)
e2
)
σ
e'
σ
[]
head_step
(
App
(
Rec
f
x
e1
)
e2
)
σ
None
e'
σ
[]
|
UnOpS
op
e
v
v'
σ
:
|
UnOpS
op
e
v
v'
σ
:
to_val
e
=
Some
v
→
to_val
e
=
Some
v
→
un_op_eval
op
v
=
Some
v'
→
un_op_eval
op
v
=
Some
v'
→
head_step
(
UnOp
op
e
)
σ
(
of_val
v'
)
σ
[]
head_step
(
UnOp
op
e
)
σ
None
(
of_val
v'
)
σ
[]
|
BinOpS
op
e1
e2
v1
v2
v'
σ
:
|
BinOpS
op
e1
e2
v1
v2
v'
σ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
bin_op_eval
op
v1
v2
=
Some
v'
→
bin_op_eval
op
v1
v2
=
Some
v'
→
head_step
(
BinOp
op
e1
e2
)
σ
(
of_val
v'
)
σ
[]
head_step
(
BinOp
op
e1
e2
)
σ
None
(
of_val
v'
)
σ
[]
|
IfTrueS
e1
e2
σ
:
|
IfTrueS
e1
e2
σ
:
head_step
(
If
(
Lit
$
LitBool
true
)
e1
e2
)
σ
e1
σ
[]
head_step
(
If
(
Lit
$
LitBool
true
)
e1
e2
)
σ
None
e1
σ
[]
|
IfFalseS
e1
e2
σ
:
|
IfFalseS
e1
e2
σ
:
head_step
(
If
(
Lit
$
LitBool
false
)
e1
e2
)
σ
e2
σ
[]
head_step
(
If
(
Lit
$
LitBool
false
)
e1
e2
)
σ
None
e2
σ
[]
|
FstS
e1
v1
e2
v2
σ
:
|
FstS
e1
v1
e2
v2
σ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
head_step
(
Fst
(
Pair
e1
e2
))
σ
e1
σ
[]
head_step
(
Fst
(
Pair
e1
e2
))
σ
None
e1
σ
[]
|
SndS
e1
v1
e2
v2
σ
:
|
SndS
e1
v1
e2
v2
σ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
head_step
(
Snd
(
Pair
e1
e2
))
σ
e2
σ
[]
head_step
(
Snd
(
Pair
e1
e2
))
σ
None
e2
σ
[]
|
CaseLS
e0
v0
e1
e2
σ
:
|
CaseLS
e0
v0
e1
e2
σ
:
to_val
e0
=
Some
v0
→
to_val
e0
=
Some
v0
→
head_step
(
Case
(
InjL
e0
)
e1
e2
)
σ
(
App
e1
e0
)
σ
[]
head_step
(
Case
(
InjL
e0
)
e1
e2
)
σ
None
(
App
e1
e0
)
σ
[]
|
CaseRS
e0
v0
e1
e2
σ
:
|
CaseRS
e0
v0
e1
e2
σ
:
to_val
e0
=
Some
v0
→
to_val
e0
=
Some
v0
→
head_step
(
Case
(
InjR
e0
)
e1
e2
)
σ
(
App
e2
e0
)
σ
[]
head_step
(
Case
(
InjR
e0
)
e1
e2
)
σ
None
(
App
e2
e0
)
σ
[]
|
ForkS
e
σ
:
|
ForkS
e
σ
:
head_step
(
Fork
e
)
σ
(
Lit
LitUnit
)
σ
[
e
]
head_step
(
Fork
e
)
σ
None
(
Lit
LitUnit
)
σ
[
e
]
|
AllocS
e
v
σ
l
:
|
AllocS
e
v
σ
l
:
to_val
e
=
Some
v
→
σ
!!
l
=
None
→
to_val
e
=
Some
v
→
σ
!!
l
=
None
→
head_step
(
Alloc
e
)
σ
(
Lit
$
LitLoc
l
)
(<[
l
:
=
v
]>
σ
)
[]
head_step
(
Alloc
e
)
σ
None
(
Lit
$
LitLoc
l
)
(<[
l
:
=
v
]>
σ
)
[]
|
LoadS
l
v
σ
:
|
LoadS
l
v
σ
:
σ
!!
l
=
Some
v
→
σ
!!
l
=
Some
v
→
head_step
(
Load
(
Lit
$
LitLoc
l
))
σ
(
of_val
v
)
σ
[]
head_step
(
Load
(
Lit
$
LitLoc
l
))
σ
None
(
of_val
v
)
σ
[]
|
StoreS
l
e
v
σ
:
|
StoreS
l
e
v
σ
:
to_val
e
=
Some
v
→
is_Some
(
σ
!!
l
)
→
to_val
e
=
Some
v
→
is_Some
(
σ
!!
l
)
→
head_step
(
Store
(
Lit
$
LitLoc
l
)
e
)
σ
(
Lit
LitUnit
)
(<[
l
:
=
v
]>
σ
)
[]
head_step
(
Store
(
Lit
$
LitLoc
l
)
e
)
σ
None
(
Lit
LitUnit
)
(<[
l
:
=
v
]>
σ
)
[]
|
CasFailS
l
e1
v1
e2
v2
vl
σ
:
|
CasFailS
l
e1
v1
e2
v2
vl
σ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
σ
!!
l
=
Some
vl
→
vl
≠
v1
→
σ
!!
l
=
Some
vl
→
vl
≠
v1
→
vals_cas_compare_safe
vl
v1
→
vals_cas_compare_safe
vl
v1
→
head_step
(
CAS
(
Lit
$
LitLoc
l
)
e1
e2
)
σ
(
Lit
$
LitBool
false
)
σ
[]
head_step
(
CAS
(
Lit
$
LitLoc
l
)
e1
e2
)
σ
None
(
Lit
$
LitBool
false
)
σ
[]
|
CasSucS
l
e1
v1
e2
v2
σ
:
|
CasSucS
l
e1
v1
e2
v2
σ
:
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
σ
!!
l
=
Some
v1
→
σ
!!
l
=
Some
v1
→
vals_cas_compare_safe
v1
v1
→
vals_cas_compare_safe
v1
v1
→
head_step
(
CAS
(
Lit
$
LitLoc
l
)
e1
e2
)
σ
(
Lit
$
LitBool
true
)
(<[
l
:
=
v2
]>
σ
)
[]
head_step
(
CAS
(
Lit
$
LitLoc
l
)
e1
e2
)
σ
None
(
Lit
$
LitBool
true
)
(<[
l
:
=
v2
]>
σ
)
[]
|
FaaS
l
i1
e2
i2
σ
:
|
FaaS
l
i1
e2
i2
σ
:
to_val
e2
=
Some
(
LitV
(
LitInt
i2
))
→
to_val
e2
=
Some
(
LitV
(
LitInt
i2
))
→
σ
!!
l
=
Some
(
LitV
(
LitInt
i1
))
→
σ
!!
l
=
Some
(
LitV
(
LitInt
i1
))
→
head_step
(
FAA
(
Lit
$
LitLoc
l
)
e2
)
σ
(
Lit
$
LitInt
i1
)
(<[
l
:
=
LitV
(
LitInt
(
i1
+
i2
))]>
σ
)
[].
head_step
(
FAA
(
Lit
$
LitLoc
l
)
e2
)
σ
None
(
Lit
$
LitInt
i1
)
(<[
l
:
=
LitV
(
LitInt
(
i1
+
i2
))]>
σ
)
[].
(** Basic properties about the language *)
(** Basic properties about the language *)
Instance
fill_item_inj
Ki
:
Inj
(=)
(=)
(
fill_item
Ki
).
Instance
fill_item_inj
Ki
:
Inj
(=)
(=)
(
fill_item
Ki
).
...
@@ -480,11 +481,11 @@ Lemma fill_item_val Ki e :
...
@@ -480,11 +481,11 @@ Lemma fill_item_val Ki e :
is_Some
(
to_val
(
fill_item
Ki
e
))
→
is_Some
(
to_val
e
).
is_Some
(
to_val
(
fill_item
Ki
e
))
→
is_Some
(
to_val
e
).
Proof
.
intros
[
v
?].
destruct
Ki
;
simplify_option_eq
;
eauto
.
Qed
.
Proof
.
intros
[
v
?].
destruct
Ki
;
simplify_option_eq
;
eauto
.
Qed
.
Lemma
val_head_stuck
e1
σ
1 e2
σ
2
efs
:
head_step
e1
σ
1 e2
σ
2
efs
→
to_val
e1
=
None
.
Lemma
val_head_stuck
e1
σ
1
κ
e2
σ
2
efs
:
head_step
e1
σ
1
κ
e2
σ
2
efs
→
to_val
e1
=
None
.
Proof
.
destruct
1
;
naive_solver
.
Qed
.
Proof
.
destruct
1
;
naive_solver
.
Qed
.
Lemma
head_ctx_step_val
Ki
e
σ
1 e2
σ
2
efs
:
Lemma
head_ctx_step_val
Ki
e
σ
1
κ
e2
σ
2
efs
:
head_step
(
fill_item
Ki
e
)
σ
1 e2
σ
2
efs
→
is_Some
(
to_val
e
).
head_step
(
fill_item
Ki
e
)
σ
1
κ
e2
σ
2
efs
→
is_Some
(
to_val
e
).
Proof
.
destruct
Ki
;
inversion_clear
1
;
simplify_option_eq
;
by
eauto
.
Qed
.
Proof
.
destruct
Ki
;
inversion_clear
1
;
simplify_option_eq
;
by
eauto
.
Qed
.
Lemma
fill_item_no_val_inj
Ki1
Ki2
e1
e2
:
Lemma
fill_item_no_val_inj
Ki1
Ki2
e1
e2
:
...
@@ -499,7 +500,7 @@ Qed.
...
@@ -499,7 +500,7 @@ Qed.
Lemma
alloc_fresh
e
v
σ
:
Lemma
alloc_fresh
e
v
σ
:
let
l
:
=
fresh
(
dom
(
gset
loc
)
σ
)
in
let
l
:
=
fresh
(
dom
(
gset
loc
)
σ
)
in
to_val
e
=
Some
v
→
head_step
(
Alloc
e
)
σ
(
Lit
(
LitLoc
l
))
(<[
l
:
=
v
]>
σ
)
[].
to_val
e
=
Some
v
→
head_step
(
Alloc
e
)
σ
None
(
Lit
(
LitLoc
l
))
(<[
l
:
=
v
]>
σ
)
[].
Proof
.
by
intros
;
apply
AllocS
,
(
not_elem_of_dom
(
D
:
=
gset
loc
)),
is_fresh
.
Qed
.
Proof
.
by
intros
;
apply
AllocS
,
(
not_elem_of_dom
(
D
:
=
gset
loc
)),
is_fresh
.
Qed
.
(* Misc *)
(* Misc *)
...
...
theories/heap_lang/lifting.v
View file @
936eeb48
...
@@ -35,7 +35,7 @@ Ltac inv_head_step :=
...
@@ -35,7 +35,7 @@ Ltac inv_head_step :=
repeat
match
goal
with
repeat
match
goal
with
|
_
=>
progress
simplify_map_eq
/=
(* simplify memory stuff *)
|
_
=>
progress
simplify_map_eq
/=
(* simplify memory stuff *)
|
H
:
to_val
_
=
Some
_
|-
_
=>
apply
of_to_val
in
H
|
H
:
to_val
_
=
Some
_
|-
_
=>
apply
of_to_val
in
H
|
H
:
head_step
?e
_
_
_
_
|-
_
=>
|
H
:
head_step
?e
_
_
_
_
_
|-
_
=>
try
(
is_var
e
;
fail
1
)
;
(* inversion yields many goals if [e] is a variable
try
(
is_var
e
;
fail
1
)
;
(* inversion yields many goals if [e] is a variable
and can thus better be avoided. *)
and can thus better be avoided. *)
inversion
H
;
subst
;
clear
H
inversion
H
;
subst
;
clear
H
...
@@ -48,7 +48,7 @@ Local Hint Constructors head_step.
...
@@ -48,7 +48,7 @@ Local Hint Constructors head_step.
Local
Hint
Resolve
alloc_fresh
.
Local
Hint
Resolve
alloc_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
4
eexists
;
econstructor
;
eauto
.
Local
Ltac
solve_exec_puredet
:
=
simpl
;
intros
;
by
inv_head_step
.
Local
Ltac
solve_exec_puredet
:
=
simpl
;
intros
;
by
inv_head_step
.
Local
Ltac
solve_pure_exec
:
=
Local
Ltac
solve_pure_exec
:
=
unfold
IntoVal
in
*
;
unfold
IntoVal
in
*
;
...
@@ -109,14 +109,14 @@ Lemma wp_fork s E e Φ :
...
@@ -109,14 +109,14 @@ Lemma wp_fork s E e Φ :
▷
WP
e
@
s
;
⊤
{{
_
,
True
}}
-
∗
▷
Φ
(
LitV
LitUnit
)
-
∗
WP
Fork
e
@
s
;
E
{{
Φ
}}.
▷
WP
e
@
s
;
⊤
{{
_
,
True
}}
-
∗
▷
Φ
(
LitV
LitUnit
)
-
∗
WP
Fork
e
@
s
;
E
{{
Φ
}}.
Proof
.
Proof
.
iIntros
"He HΦ"
.
iIntros
"He HΦ"
.
iApply
wp_lift_pure_det_head_step
;
[
auto
|
intros
;
inv_head_step
;
eauto
|].
iApply
wp_lift_pure_det_head_step
;
[
e
auto
|
intros
;
inv_head_step
;
eauto
|].
iModIntro
;
iNext
;
iIntros
"!> /= {$He}"
.
by
iApply
wp_value
.
iModIntro
;
iNext
;
iIntros
"!> /= {$He}"
.
by
iApply
wp_value
.
Qed
.
Qed
.
Lemma
twp_fork
s
E
e
Φ
:
Lemma
twp_fork
s
E
e
Φ
:
WP
e
@
s
;
⊤
[{
_
,
True
}]
-
∗
Φ
(
LitV
LitUnit
)
-
∗
WP
Fork
e
@
s
;
E
[{
Φ
}].
WP
e
@
s
;
⊤
[{
_
,
True
}]
-
∗
Φ
(
LitV
LitUnit
)
-
∗
WP
Fork
e
@
s
;
E
[{
Φ
}].
Proof
.
Proof
.
iIntros
"He HΦ"
.
iIntros
"He HΦ"
.
iApply
twp_lift_pure_det_head_step
;
[
auto
|
intros
;
inv_head_step
;
eauto
|].
iApply
twp_lift_pure_det_head_step
;
[
e
auto
|
intros
;
inv_head_step
;
eauto
|].
iIntros
"!> /= {$He}"
.
by
iApply
twp_value
.
iIntros
"!> /= {$He}"
.
by
iApply
twp_value
.
Qed
.
Qed
.
...
@@ -126,8 +126,8 @@ Lemma wp_alloc s E e v :
...
@@ -126,8 +126,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
)
"Hσ !>"
;
iSplit
;
first
by
auto
.
iIntros
(
σ
1
)
"Hσ !>"
;
iSplit
;
first
by
e
auto
.
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iNext
;
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
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -136,8 +136,8 @@ Lemma twp_alloc s E e v :
...
@@ -136,8 +136,8 @@ 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
)
"Hσ !>"
;
iSplit
;
first
by
auto
.
iIntros
(
σ
1
)
"Hσ !>"
;
iSplit
;
first
by
e
auto
.
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
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -148,7 +148,7 @@ Proof.
...
@@ -148,7 +148,7 @@ 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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
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
(
κ
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
:
...
@@ -157,7 +157,7 @@ Proof.
...
@@ -157,7 +157,7 @@ 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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
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
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -168,7 +168,7 @@ Proof.
...
@@ -168,7 +168,7 @@ 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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
eauto
6
.
iNext
;
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
=>//.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -179,7 +179,7 @@ Proof.
...
@@ -179,7 +179,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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iIntros
(
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
eauto
6
.
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
=>//.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -192,7 +192,7 @@ Proof.
...
@@ -192,7 +192,7 @@ 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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
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
:
...
@@ -203,7 +203,7 @@ Proof.
...
@@ -203,7 +203,7 @@ Proof.
iIntros
(<-
[
v2
<-]
??
Φ
)
"Hl HΦ"
.
iIntros
(<-
[
v2
<-]
??
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
eauto
.
iIntros
(
κ
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -215,7 +215,7 @@ Proof.
...
@@ -215,7 +215,7 @@ 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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
eauto
.
iNext
;
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
=>//.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -227,7 +227,7 @@ Proof.
...
@@ -227,7 +227,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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
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
=>//.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -240,7 +240,7 @@ Proof.
...
@@ -240,7 +240,7 @@ 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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
eauto
.
iNext
;
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
=>//.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
@@ -252,7 +252,7 @@ Proof.
...
@@ -252,7 +252,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
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iIntros
(
σ
1
)
"Hσ !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
iSplit
;
first
by
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
=>//.
by
iApply
"HΦ"
.
iModIntro
.
iSplit
=>//.
by
iApply
"HΦ"
.
Qed
.
Qed
.
...
...
theories/heap_lang/total_adequacy.v
View file @
936eeb48
...
@@ -6,7 +6,7 @@ Set Default Proof Using "Type".
...
@@ -6,7 +6,7 @@ Set Default Proof Using "Type".
Definition
heap_total
Σ
`
{
heapPreG
Σ
}
s
e
σ
φ
:
Definition
heap_total
Σ
`
{
heapPreG
Σ
}
s
e
σ
φ
:
(
∀
`
{
heapG
Σ
},
WP
e
@
s
;
⊤
[{
v
,
⌜φ
v
⌝
}]%
I
)
→
(
∀
`
{
heapG
Σ
},
WP
e
@
s
;
⊤
[{
v
,
⌜φ
v
⌝
}]%
I
)
→
sn
step
([
e
],
σ
).
sn
erased_
step
([
e
],
σ
).
Proof
.
Proof
.
intros
Hwp
;
eapply
(
twp_total
_
_
)
;
iIntros
(?)
""
.
intros
Hwp
;
eapply
(
twp_total
_
_
)
;
iIntros
(?)
""
.
iMod
(
gen_heap_init
σ
)
as
(?)
"Hh"
.
iMod
(
gen_heap_init
σ
)
as
(?)
"Hh"
.
...
...
theories/program_logic/adequacy.v
View file @
936eeb48
...
@@ -37,26 +37,26 @@ Qed.
...
@@ -37,26 +37,26 @@ Qed.
Record
adequate
{
Λ
}
(
s
:
stuckness
)
(
e1
:
expr
Λ
)
(
σ
1
:
state
Λ
)
Record
adequate
{
Λ
}
(
s
:
stuckness
)
(
e1
:
expr
Λ
)
(
σ
1
:
state
Λ
)
(
φ
:
val
Λ
→
state
Λ
→
Prop
)
:
=
{
(
φ
:
val
Λ
→
state
Λ
→
Prop
)
:
=
{
adequate_result
t2
σ
2
v2
:
adequate_result
t2
σ
2
v2
:
rtc
step
([
e1
],
σ
1
)
(
of_val
v2
::
t2
,
σ
2
)
→
φ
v2
σ
2
;
rtc
erased_
step
([
e1
],
σ
1
)
(
of_val
v2
::
t2
,
σ
2
)
→
φ
v2
σ
2
;
adequate_not_stuck
t2
σ
2 e2
:
adequate_not_stuck
t2
σ
2 e2
:
s
=
NotStuck
→
s
=
NotStuck
→
rtc
step
([
e1
],
σ
1
)
(
t2
,
σ
2
)
→
rtc
erased_
step
([
e1
],
σ
1
)
(
t2
,
σ
2
)
→
e2
∈
t2
→
(
is_Some
(
to_val
e2
)
∨
reducible
e2
σ
2
)
e2
∈
t2
→
(
is_Some
(
to_val
e2
)
∨
reducible
e2
σ
2
)
}.
}.
Theorem
adequate_tp_safe
{
Λ
}
(
e1
:
expr
Λ
)
t2
σ
1
σ
2
φ
:
Theorem
adequate_tp_safe
{
Λ
}
(
e1
:
expr
Λ
)
t2
σ
1
σ
2
φ
:
adequate
NotStuck
e1
σ
1
φ
→
adequate
NotStuck
e1
σ
1
φ
→
rtc
step
([
e1
],
σ
1
)
(
t2
,
σ
2
)
→
rtc
erased_
step
([
e1
],
σ
1
)
(
t2
,
σ
2
)
→
Forall
(
λ
e
,
is_Some
(
to_val
e
))
t2
∨
∃
t3
σ
3
,
step
(
t2
,
σ
2
)
(
t3
,
σ
3
).
Forall
(
λ
e
,
is_Some
(
to_val
e
))
t2
∨
∃
κ
t3
σ
3
,
step
(
t2
,
σ
2
)
κ
(
t3
,
σ
3
).
Proof
.
Proof
.
intros
Had
?.
intros
Had
?.
destruct
(
decide
(
Forall
(
λ
e
,
is_Some
(
to_val
e
))
t2
))
as
[|
Ht2
]
;
[
by
left
|].
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
).
apply
(
not_Forall_Exists
_
),
Exists_exists
in
Ht2
;
destruct
Ht2
as
(
e2
&?&
He2
).
destruct
(
adequate_not_stuck
NotStuck
e1
σ
1
φ
Had
t2
σ
2 e2
)
as
[?|(
e3
&
σ
3
&
efs
&?)]
;
destruct
(
adequate_not_stuck
NotStuck
e1
σ
1
φ
Had
t2
σ
2 e2
)
as
[?|(
κ
&
e3
&
σ
3
&
efs
&?)]
;
rewrite
?eq_None_not_Some
;
auto
.
rewrite
?eq_None_not_Some
;
auto
.
{
exfalso
.
eauto
.
}
{
exfalso
.
eauto
.
}
destruct
(
elem_of_list_split
t2
e2
)
as
(
t2'
&
t2''
&->)
;
auto
.
destruct
(
elem_of_list_split
t2
e2
)
as
(
t2'
&
t2''
&->)
;
auto
.
right
;
exists
(
t2'
++
e3
::
t2''
++
efs
),
σ
3
;
econstructor
;
eauto
.
right
;
exists
κ
,
(
t2'
++
e3
::
t2''
++
efs
),
σ
3
;
econstructor
;
eauto
.
Qed
.
Qed
.
Section
adequacy
.
Section
adequacy
.
...
@@ -71,39 +71,41 @@ Notation world σ := (world' ⊤ σ) (only parsing).
...
@@ -71,39 +71,41 @@ Notation world σ := (world' ⊤ σ) (only parsing).
Notation
wptp
s
t
:
=
([
∗
list
]
ef
∈
t
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})%
I
.
Notation
wptp
s
t
:
=
([
∗
list
]
ef
∈
t
,
WP
ef
@
s
;
⊤
{{
_
,
True
}})%
I
.
Lemma
wp_step
s
E
e1
σ
1 e2
σ
2
efs
Φ
:
Lemma
wp_step
s
E
e1
σ
1
κ
e2
σ
2
efs
Φ
:
prim_step
e1
σ
1 e2
σ
2
efs
→
prim_step
e1
σ
1
κ
e2
σ
2
efs
→
world'
E
σ
1
∗
WP
e1
@
s
;
E
{{
Φ
}}
world'
E
σ
1
∗
WP
e1
@
s
;
E
{{
Φ
}}
==
∗
▷
|==>
◇
(
world'
E
σ
2
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
wptp
s
efs
).
==
∗
▷
|==>
◇
(
world'
E
σ
2
∗
WP
e2
@
s
;
E
{{
Φ
}}
∗
wptp
s
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"
$!
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
.
Lemma
wptp_step
s
e1
t1
t2
σ
1
σ
2
Φ
:
(* should we be able to say that κs = κ :: κs'? *)
step
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
Lemma
wptp_step
s
e1
t1
t2
κ
σ
1
σ
2
Φ
:
step
(
e1
::
t1
,
σ
1
)
κ
(
t2
,
σ
2
)
→
world
σ
1
∗
WP
e1
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t1
world
σ
1
∗
WP
e1
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t1
==
∗
∃
e2
t2'
,
⌜
t2
=
e2
::
t2'
⌝
∗
▷
|==>
◇
(
world
σ
2
∗
WP
e2
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t2'
).
==
∗
∃
e2
t2'
,
⌜
t2
=
e2
::
t2'
⌝
∗
▷
|==>
◇
(
world
σ
2
∗
WP
e2
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t2'
).
Proof
.
Proof
.
iIntros
(
Hstep
)
"(HW & He & Ht)"
.
iIntros
(
Hstep
)
"(HW & He & Ht)"
.
destruct
Hstep
as
[
e1'
σ
1
'
e2'
σ
2
'
efs
[|?
t1'
]
t2'
??
Hstep
]
;
simplify_eq
/=.
destruct
Hstep
as
[
e1'
σ
1
'
e2'
σ
2
'
efs
[|?
t1'
]
t2'
??
Hstep
]
;
simplify_eq
/=.
-
iExists
e2'
,
(
t2'
++
efs
)
;
iSplitR
;
first
eauto
.
-
iExists
e2'
,
(
t2'
++
efs
)
;
iSplitR
;
first
by
eauto
.
iFrame
"Ht"
.
iApply
wp_step
;
eauto
with
iFrame
.
iFrame
"Ht"
.
iApply
wp_step
;
eauto
with
iFrame
.
-
iExists
e
,
(
t1'
++
e2'
::
t2'
++
efs
)
;
iSplitR
;
first
eauto
.
-
iExists
e
,
(
t1'
++
e2'
::
t2'
++
efs
)
;
iSplitR
;
first
eauto
.
iDestruct
"Ht"
as
"($ & He' & $)"
.
iFrame
"He"
.
iDestruct
"Ht"
as
"($ & He' & $)"
.
iFrame
"He"
.
iApply
wp_step
;
eauto
with
iFrame
.
iApply
wp_step
;
eauto
with
iFrame
.
Qed
.
Qed
.
Lemma
wptp_steps
s
n
e1
t1
t2
σ
1
σ
2
Φ
:
Lemma
wptp_steps
s
n
e1
t1
κ
s
t2
σ
1
σ
2
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
nsteps
n
(
e1
::
t1
,
σ
1
)
κ
s
(
t2
,
σ
2
)
→
world
σ
1
∗
WP
e1
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t1
⊢
world
σ
1
∗
WP
e1
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t1
⊢
Nat
.
iter
(
S
n
)
(
λ
P
,
|==>
▷
P
)
(
∃
e2
t2'
,
Nat
.
iter
(
S
n
)
(
λ
P
,
|==>
▷
P
)
(
∃
e2
t2'
,
⌜
t2
=
e2
::
t2'
⌝
∗
world
σ
2
∗
WP
e2
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t2'
).
⌜
t2
=
e2
::
t2'
⌝
∗
world
σ
2
∗
WP
e2
@
s
;
⊤
{{
Φ
}}
∗
wptp
s
t2'
).
Proof
.
Proof
.
revert
e1
t1
t2
σ
1
σ
2
;
simpl
;
induction
n
as
[|
n
IH
]=>