Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Dan Frumin
ReLoC-v1
Commits
9a182da8
Commit
9a182da8
authored
Jul 02, 2016
by
Robbert Krebbers
Browse files
Make some names more consistent.
parent
88c9a1e0
Changes
12
Hide whitespace changes
Inline
Side-by-side
F_mu_ref_par/context_refinement.v
View file @
9a182da8
From
iris_logrel
.
F_mu_ref_par
Require
Export
fundamental_binary
.
Inductive
ctx_item
:=
|
CTX_
Lam
|
CTX_
Rec
|
CTX_AppL
(
e2
:
expr
)
|
CTX_AppR
(
e1
:
expr
)
(
*
Products
*
)
...
...
@@ -42,7 +42,7 @@ Inductive ctx_item :=
Fixpoint
fill_ctx_item
(
ctx
:
ctx_item
)
(
e
:
expr
)
:
expr
:=
match
ctx
with
|
CTX_
Lam
=>
Lam
e
|
CTX_
Rec
=>
Rec
e
|
CTX_AppL
e2
=>
App
e
e2
|
CTX_AppR
e1
=>
App
e1
e
|
CTX_PairL
e2
=>
Pair
e
e2
...
...
@@ -80,8 +80,8 @@ Definition fill_ctx (K : ctx) (e : expr) : expr := foldr fill_ctx_item e K.
(
**
typed
ctx
*
)
Inductive
typed_ctx_item
:
ctx_item
→
list
type
→
type
→
list
type
→
type
→
Prop
:=
|
TP_CTX_
Lam
Γ
τ
τ'
:
typed_ctx_item
CTX_
Lam
(
TArrow
τ
τ'
::
τ
::
Γ
)
τ'
Γ
(
TArrow
τ
τ'
)
|
TP_CTX_
Rec
Γ
τ
τ'
:
typed_ctx_item
CTX_
Rec
(
TArrow
τ
τ'
::
τ
::
Γ
)
τ'
Γ
(
TArrow
τ
τ'
)
|
TP_CTX_AppL
Γ
e2
τ
τ'
:
typed
Γ
e2
τ
→
typed_ctx_item
(
CTX_AppL
e2
)
Γ
(
TArrow
τ
τ'
)
Γ
τ'
...
...
@@ -211,22 +211,22 @@ Section bin_log_related_under_typed_ctx.
-
inversion_clear
1
as
[
|?
?
?
?
?
?
?
?
Hx1
Hx2
].
intros
H3
Δ
H
Δ
.
specialize
(
IHK
_
_
_
_
e
e
'
H1
H2
Hx2
H3
).
inversion
Hx1
;
subst
;
simpl
.
+
eapply
bin_log_related_
Lam
;
eauto
;
+
eapply
bin_log_related_
rec
;
eauto
;
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_ctx_n_closed
_
_
_
_
_
_
_
H
)
end
.
+
eapply
bin_log_related_
A
pp
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
A
pp
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
P
air
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
P
air
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
F
st
;
eauto
.
+
eapply
bin_log_related_
S
nd
;
eauto
.
+
eapply
bin_log_related_
I
nj
L
;
eauto
.
+
eapply
bin_log_related_
I
nj
R
;
eauto
.
+
eapply
bin_log_related_
a
pp
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
a
pp
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
p
air
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
p
air
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
f
st
;
eauto
.
+
eapply
bin_log_related_
s
nd
;
eauto
.
+
eapply
bin_log_related_
i
nj
l
;
eauto
.
+
eapply
bin_log_related_
i
nj
r
;
eauto
.
+
match
goal
with
H
:
typed_ctx_item
_
_
_
_
_
|-
_
=>
inversion
H
;
subst
end
.
eapply
bin_log_related_
C
ase
;
eapply
bin_log_related_
c
ase
;
eauto
using
binary_fundamental
;
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_n_closed
_
_
_
H
)
...
...
@@ -234,7 +234,7 @@ Section bin_log_related_under_typed_ctx.
+
match
goal
with
H
:
typed_ctx_item
_
_
_
_
_
|-
_
=>
inversion
H
;
subst
end
.
eapply
bin_log_related_
C
ase
;
eapply
bin_log_related_
c
ase
;
eauto
using
binary_fundamental
;
try
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_n_closed
_
_
_
H
)
...
...
@@ -245,7 +245,7 @@ Section bin_log_related_under_typed_ctx.
+
match
goal
with
H
:
typed_ctx_item
_
_
_
_
_
|-
_
=>
inversion
H
;
subst
end
.
eapply
bin_log_related_
C
ase
;
eapply
bin_log_related_
c
ase
;
eauto
using
binary_fundamental
;
try
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_n_closed
_
_
_
H
)
...
...
@@ -253,22 +253,22 @@ Section bin_log_related_under_typed_ctx.
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_ctx_n_closed
_
_
_
_
_
_
_
H
)
end
.
+
eapply
bin_log_related_
I
f
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_
I
f
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_
I
f
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_nat_bin
_
op
;
+
eapply
bin_log_related_
i
f
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_
i
f
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_
i
f
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_nat_binop
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_nat_bin
_
op
;
+
eapply
bin_log_related_nat_binop
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_
F
old
;
eauto
.
+
eapply
bin_log_related_
U
nfold
;
eauto
.
+
eapply
bin_log_related_
TL
am
;
eauto
with
typeclass_instances
.
+
eapply
bin_log_related_
TA
pp
;
eauto
.
+
eapply
bin_log_related_
F
ork
;
eauto
.
+
eapply
bin_log_related_
A
lloc
;
eauto
.
+
eapply
bin_log_related_
L
oad
;
eauto
.
+
eapply
bin_log_related_
S
tore
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
S
tore
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
f
old
;
eauto
.
+
eapply
bin_log_related_
u
nfold
;
eauto
.
+
eapply
bin_log_related_
tl
am
;
eauto
with
typeclass_instances
.
+
eapply
bin_log_related_
ta
pp
;
eauto
.
+
eapply
bin_log_related_
f
ork
;
eauto
.
+
eapply
bin_log_related_
a
lloc
;
eauto
.
+
eapply
bin_log_related_
l
oad
;
eauto
.
+
eapply
bin_log_related_
s
tore
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_
s
tore
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_CAS
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_CAS
;
eauto
using
binary_fundamental
.
+
eapply
bin_log_related_CAS
;
eauto
using
binary_fundamental
.
...
...
F_mu_ref_par/examples/counter.v
View file @
9a182da8
...
...
@@ -3,27 +3,27 @@ From iris_logrel.F_mu_ref_par Require Export examples.lock.
From
iris_logrel
.
F_mu_ref_par
Require
Import
soundness_binary
.
Definition
CG_increment
(
x
:
expr
)
:
expr
:=
Lam
(
Store
x
.[
ren
(
+
2
)]
(
BinOp
Add
(
♯
1
)
(
Load
x
.[
ren
(
+
2
)]))).
Rec
(
Store
x
.[
ren
(
+
2
)]
(
BinOp
Add
(
♯
1
)
(
Load
x
.[
ren
(
+
2
)]))).
Definition
CG_locked_increment
(
x
l
:
expr
)
:
expr
:=
with_lock
(
CG_increment
x
)
l
.
Definition
CG_locked_incrementV
(
x
l
:
expr
)
:
val
:=
with_lockV
(
CG_increment
x
)
l
.
Definition
counter_read
(
x
:
expr
)
:
expr
:=
Lam
(
Load
x
.[
ren
(
+
2
)]).
Definition
counter_readV
(
x
:
expr
)
:
val
:=
Lam
V
(
Load
x
.[
ren
(
+
2
)]).
Definition
counter_read
(
x
:
expr
)
:
expr
:=
Rec
(
Load
x
.[
ren
(
+
2
)]).
Definition
counter_readV
(
x
:
expr
)
:
val
:=
Rec
V
(
Load
x
.[
ren
(
+
2
)]).
Definition
CG_counter_body
(
x
l
:
expr
)
:
expr
:=
Pair
(
CG_locked_increment
x
l
)
(
counter_read
x
).
Definition
CG_counter
:
expr
:=
App
(
Lam
$
App
(
Lam
(
CG_counter_body
(
Var
1
)
(
Var
3
)))
(
Rec
$
App
(
Rec
(
CG_counter_body
(
Var
1
)
(
Var
3
)))
(
Alloc
(
♯
0
)))
newlock
.
Definition
FG_increment
(
x
:
expr
)
:
expr
:=
Lam
$
App
(
Lam
$
Rec
$
App
(
Rec
$
(
*
try
increment
*
)
If
(
CAS
x
.[
ren
(
+
4
)]
(
Var
1
)
(
BinOp
Add
(
♯
1
)
(
Var
1
)))
Unit
(
*
increment
succeeds
we
return
unit
*
)
...
...
@@ -32,7 +32,7 @@ Definition FG_increment (x : expr) : expr :=
Definition
FG_counter_body
(
x
:
expr
)
:
expr
:=
Pair
(
FG_increment
x
)
(
counter_read
x
).
Definition
FG_counter
:
expr
:=
App
(
Lam
(
FG_counter_body
(
Var
1
)))
(
Alloc
(
♯
0
)).
App
(
Rec
(
FG_counter_body
(
Var
1
)))
(
Alloc
(
♯
0
)).
Section
CG_Counter
.
Context
`
{
cfgSG
Σ
,
heapIG
Σ
}
.
...
...
@@ -63,13 +63,13 @@ Section CG_Counter.
={
E
}=>
j
⤇
fill
K
(
Unit
)
★
x
↦ₛ
(
♯
v
(
S
n
)).
Proof
.
iIntros
{
HNE
}
"[#Hspec [Hx Hj]]"
.
unfold
CG_increment
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_load
_
_
j
(
K
++
[
StoreRCtx
(
LocV
_
);
BinOpRCtx
_
(
♯
v
_
)])
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
rewrite
?
fill_app
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
rewrite
?
fill_app
.
simpl
.
iPvs
(
step_nat_bin
_
op
_
_
j
(
K
++
[
StoreRCtx
(
LocV
_
)])
iPvs
(
step_nat_binop
_
_
j
(
K
++
[
StoreRCtx
(
LocV
_
)])
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
rewrite
?
fill_app
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
simpl
.
...
...
@@ -164,7 +164,7 @@ Section CG_Counter.
={
E
}=>
j
⤇
fill
K
(
♯
n
)
★
x
↦ₛ
(
♯
v
n
).
Proof
.
intros
HNE
.
iIntros
"[#Hspec [Hx Hj]]"
.
unfold
counter_read
.
iPvs
(
step_
lam
_
_
j
K
_
Unit
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
Unit
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_load
_
_
j
K
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
{
by
iFrame
"Hspec Hj"
.
}
...
...
@@ -260,24 +260,24 @@ Section CG_Counter.
{
iDestruct
(
interp_env_length
with
"HΓ"
)
as
%
[
=
].
}
iClear
"HΓ"
.
cbn
-
[
FG_counter
CG_counter
].
rewrite
?
empty_env_subst
/
CG_counter
/
FG_counter
.
iPvs
(
steps_newlock
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
)])
_
with
"[Hj]"
)
iPvs
(
steps_newlock
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
)])
_
with
"[Hj]"
)
as
{
l
}
"[Hj Hl]"
;
eauto
.
{
rewrite
fill_app
/=
.
by
iFrame
.
}
rewrite
fill_app
/=
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
rewrite
CG_locked_increment_subst
/=
.
rewrite
counter_read_subst
/=
.
iPvs
(
step_alloc
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
)])
_
_
_
_
with
"[Hj]"
)
iPvs
(
step_alloc
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
)])
_
_
_
_
with
"[Hj]"
)
as
{
cnt
'
}
"[Hj Hcnt']"
;
eauto
.
{
rewrite
fill_app
;
simpl
.
by
iFrame
.
}
rewrite
fill_app
;
simpl
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
rewrite
CG_locked_increment_subst
/=
.
rewrite
counter_read_subst
/=
.
Unshelve
.
all:
try
match
goal
with
|-
to_val
_
=
_
=>
auto
using
to_of_val
end
.
all:
trivial
.
iApply
(
@
wp_bind
_
_
_
[
AppRCtx
(
Lam
V
_
)]);
iApply
(
@
wp_bind
_
_
_
[
AppRCtx
(
Rec
V
_
)]);
iApply
wp_wand_l
;
iSplitR
;
[
iIntros
{
v
}
"Hv"
;
iExact
"Hv"
|
].
iApply
wp_alloc
;
trivial
;
iFrame
"Hheap"
;
iNext
;
iIntros
{
cnt
}
"Hcnt /="
.
iPvsIntro
.
...
...
@@ -288,7 +288,7 @@ Section CG_Counter.
iPvs
(
inv_alloc
counterN
with
"[Hinv]"
)
as
"#Hinv"
;
trivial
.
{
iNext
;
iExact
"Hinv"
.
}
(
*
splitting
increment
and
read
*
)
iApply
wp_
lam
;
trivial
.
iNext
.
asimpl
.
iApply
wp_
rec
;
trivial
.
iNext
.
asimpl
.
rewrite
counter_read_subst
/=
.
iApply
wp_value
;
simpl
.
{
by
rewrite
counter_read_to_val
.
}
...
...
@@ -301,19 +301,19 @@ Section CG_Counter.
rewrite
CG_locked_increment_of_val
/=
.
destruct
v
;
iDestruct
"Heq"
as
"[% %]"
;
simplify_eq
/=
.
iL
ö
b
as
"Hlat"
.
iApply
wp_
lam
;
trivial
.
asimpl
.
iNext
.
iApply
wp_
rec
;
trivial
.
asimpl
.
iNext
.
(
*
fine
-
grained
reads
the
counter
*
)
iApply
(
@
wp_bind
_
_
_
[
AppRCtx
(
Lam
V
_
)]);
iApply
(
@
wp_bind
_
_
_
[
AppRCtx
(
Rec
V
_
)]);
iApply
wp_wand_l
;
iSplitR
;
[
iIntros
{
v
}
"Hv"
;
iExact
"Hv"
|
].
iInv
>
counterN
as
{
n
}
"[Hl [Hcnt Hcnt']]"
.
iApply
wp_load
;
[
|
iFrame
"Hheap"
].
solve_ndisj
.
iIntros
"> {$Hcnt} Hcnt"
.
iPvsIntro
.
iSplitL
"Hl Hcnt Hcnt'"
;
[
iExists
_
;
iFrame
"Hl Hcnt Hcnt'"
;
trivial
|
].
iApply
wp_
lam
;
trivial
.
asimpl
.
iNext
.
iApply
wp_
rec
;
trivial
.
asimpl
.
iNext
.
(
*
fine
-
grained
performs
increment
*
)
iApply
(
@
wp_bind
_
_
_
[
IfCtx
_
_
;
CasRCtx
(
LocV
_
)
(
NatV
_
)]);
iApply
wp_wand_l
;
iSplitR
;
[
iIntros
{
v
}
"Hv"
;
iExact
"Hv"
|
].
iApply
wp_nat_bin
_
op
;
simpl
.
iApply
wp_nat_binop
;
simpl
.
iNext
.
iPvsIntro
.
iApply
(
@
wp_bind
_
_
_
[
IfCtx
_
_
]);
iApply
wp_wand_l
;
iSplitR
;
[
iIntros
{
v
}
"Hv"
;
iExact
"Hv"
|
].
...
...
@@ -343,7 +343,7 @@ Section CG_Counter.
iDestruct
"Heq"
as
"[% %]"
;
destruct
v
;
simplify_eq
/=
.
Transparent
counter_read
.
unfold
counter_read
at
2.
iApply
wp_
lam
;
trivial
.
simpl
.
iApply
wp_
rec
;
trivial
.
simpl
.
iNext
.
iInv
>
counterN
as
{
n
}
"[Hl [Hcnt Hcnt']]"
.
iPvs
(
steps_counter_read
with
"[Hj Hcnt']"
)
as
"[Hj Hcnt']"
.
{
solve_ndisj
.
}
...
...
F_mu_ref_par/examples/lock.v
View file @
9a182da8
...
...
@@ -3,19 +3,19 @@ From iris_logrel.F_mu_ref_par Require Export rules_binary typing.
Definition
newlock
:
expr
:=
Alloc
(
♭
false
).
Definition
acquire
:
expr
:=
Lam
(
If
(
CAS
(
Var
1
)
(
♭
false
)
(
♭
true
))
(
Unit
)
(
App
(
Var
0
)
(
Var
1
))).
Definition
release
:
expr
:=
Lam
(
Store
(
Var
1
)
(
♭
false
)).
Rec
(
If
(
CAS
(
Var
1
)
(
♭
false
)
(
♭
true
))
(
Unit
)
(
App
(
Var
0
)
(
Var
1
))).
Definition
release
:
expr
:=
Rec
(
Store
(
Var
1
)
(
♭
false
)).
Definition
with_lock
(
e
:
expr
)
(
l
:
expr
)
:
expr
:=
Lam
(
App
(
Lam
(
App
(
Lam
(
App
(
Lam
(
Var
3
))
(
App
release
l
.[
ren
(
+
6
)])))
Rec
(
App
(
Rec
(
App
(
Rec
(
App
(
Rec
(
Var
3
))
(
App
release
l
.[
ren
(
+
6
)])))
(
App
e
.[
ren
(
+
4
)]
(
Var
3
))))
(
App
acquire
l
.[
ren
(
+
2
)])
).
Definition
with_lockV
(
e
l
:
expr
)
:
val
:=
Lam
V
(
App
(
Lam
(
App
(
Lam
(
App
(
Lam
(
Var
3
))
(
App
release
l
.[
ren
(
+
6
)])))
Rec
V
(
App
(
Rec
(
App
(
Rec
(
App
(
Rec
(
Var
3
))
(
App
release
l
.[
ren
(
+
6
)])))
(
App
e
.[
ren
(
+
4
)]
(
Var
3
))))
(
App
acquire
l
.[
ren
(
+
2
)])
).
...
...
@@ -93,7 +93,7 @@ Section proof.
={
E
}=>
j
⤇
fill
K
Unit
★
l
↦ₛ
(
♭
v
true
).
Proof
.
iIntros
{
HNE
}
"[#Hspec [Hl Hj]]"
.
unfold
acquire
.
iPvs
(
step_
lam
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
.
done
.
iPvs
(
step_
rec
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
.
done
.
iPvs
(
step_cas_suc
_
_
j
(
K
++
[
IfCtx
_
_
])
_
_
_
_
_
_
_
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
trivial
.
rewrite
?
fill_app
.
simpl
.
...
...
@@ -114,7 +114,7 @@ Section proof.
={
E
}=>
j
⤇
fill
K
Unit
★
l
↦ₛ
(
♭
v
false
).
Proof
.
iIntros
{
HNE
}
"[#Hspec [Hl Hj]]"
.
unfold
release
.
iPvs
(
step_
lam
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
;
try
done
.
iPvs
(
step_
rec
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
;
try
done
.
iPvs
(
step_store
_
_
j
K
_
_
_
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
iFrame
"Hspec Hj"
;
trivial
.
iPvsIntro
.
iFrame
"Hj Hl"
;
trivial
.
...
...
@@ -133,26 +133,26 @@ Section proof.
={
E
}=>
j
⤇
fill
K
(#
v
)
★
Q
★
l
↦ₛ
(
♭
v
false
).
Proof
.
iIntros
{
HNE
H1
H2
}
"[#Hspec [HP [Hl Hj]]]"
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
rewrite
H1
.
iPvs
(
steps_acquire
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
)])
iPvs
(
steps_acquire
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
)])
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
rewrite
fill_app
;
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
rewrite
fill_app
;
simpl
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
rewrite
H1
.
iPvs
(
H2
(
K
++
[
AppRCtx
(
Lam
V
_
)])
with
"[Hj HP]"
)
as
"[Hj HQ]"
;
eauto
.
iPvs
(
H2
(
K
++
[
AppRCtx
(
Rec
V
_
)])
with
"[Hj HP]"
)
as
"[Hj HQ]"
;
eauto
.
rewrite
?
fill_app
/=
.
iFrame
"Hspec HP"
;
trivial
.
rewrite
?
fill_app
/=
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
steps_release
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
)])
_
_
with
"[Hj Hl]"
)
iPvs
(
steps_release
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
)])
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
rewrite
?
fill_app
/=
.
by
iFrame
.
rewrite
?
fill_app
/=
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvsIntro
;
by
iFrame
.
Unshelve
.
all:
try
match
goal
with
|-
to_val
_
=
_
=>
auto
using
to_of_val
end
.
...
...
F_mu_ref_par/examples/stack/CG_stack.v
View file @
9a182da8
...
...
@@ -7,17 +7,17 @@ Definition CG_StackType τ :=
(
*
Coarse
-
grained
push
*
)
Definition
CG_push
(
st
:
expr
)
:
expr
:=
Lam
(
Store
Rec
(
Store
(
st
.[
ren
(
+
2
)])
(
Fold
(
InjR
(
Pair
(
Var
1
)
(
Load
st
.[
ren
(
+
2
)]))))).
Definition
CG_locked_push
(
st
l
:
expr
)
:=
with_lock
(
CG_push
st
)
l
.
Definition
CG_locked_pushV
(
st
l
:
expr
)
:
val
:=
with_lockV
(
CG_push
st
)
l
.
Definition
CG_pop
(
st
:
expr
)
:
expr
:=
Lam
(
Case
(
Unfold
(
Load
st
.[
ren
(
+
2
)]))
Rec
(
Case
(
Unfold
(
Load
st
.[
ren
(
+
2
)]))
(
InjL
Unit
)
(
App
(
Lam
(
InjR
(
Fst
(
Var
2
))))
App
(
Rec
(
InjR
(
Fst
(
Var
2
))))
(
Store
st
.[
ren
(
+
3
)]
(
Snd
(
Var
0
)))
)
).
...
...
@@ -25,35 +25,35 @@ Definition CG_pop (st : expr) : expr :=
Definition
CG_locked_pop
(
st
l
:
expr
)
:=
with_lock
(
CG_pop
st
)
l
.
Definition
CG_locked_popV
(
st
l
:
expr
)
:
val
:=
with_lockV
(
CG_pop
st
)
l
.
Definition
CG_snap
(
st
l
:
expr
)
:=
with_lock
(
Lam
(
Load
st
.[
ren
(
+
2
)]))
l
.
Definition
CG_snapV
(
st
l
:
expr
)
:
val
:=
with_lockV
(
Lam
(
Load
st
.[
ren
(
+
2
)]))
l
.
Definition
CG_snap
(
st
l
:
expr
)
:=
with_lock
(
Rec
(
Load
st
.[
ren
(
+
2
)]))
l
.
Definition
CG_snapV
(
st
l
:
expr
)
:
val
:=
with_lockV
(
Rec
(
Load
st
.[
ren
(
+
2
)]))
l
.
Definition
CG_iter
(
f
:
expr
)
:
expr
:=
Lam
(
Case
(
Unfold
(
Var
1
))
Rec
(
Case
(
Unfold
(
Var
1
))
Unit
(
App
(
Lam
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
App
f
.[
ren
(
+
3
)]
(
Fst
(
Var
0
)))
)
).
Definition
CG_iterV
(
f
:
expr
)
:
val
:=
Lam
V
(
Case
(
Unfold
(
Var
1
))
Rec
V
(
Case
(
Unfold
(
Var
1
))
Unit
(
App
(
Lam
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
App
f
.[
ren
(
+
3
)]
(
Fst
(
Var
0
)))
)
).
Definition
CG_snap_iter
(
st
l
:
expr
)
:
expr
:=
Lam
(
App
(
CG_iter
(
Var
1
))
(
App
(
CG_snap
st
.[
ren
(
+
2
)]
l
.[
ren
(
+
2
)])
Unit
)).
Rec
(
App
(
CG_iter
(
Var
1
))
(
App
(
CG_snap
st
.[
ren
(
+
2
)]
l
.[
ren
(
+
2
)])
Unit
)).
Definition
CG_stack_body
(
st
l
:
expr
)
:
expr
:=
Pair
(
Pair
(
CG_locked_push
st
l
)
(
CG_locked_pop
st
l
))
(
CG_snap_iter
st
l
).
Definition
CG_stack
:
expr
:=
TLam
(
App
(
Lam
(
App
(
Lam
(
CG_stack_body
(
Var
1
)
(
Var
3
)))
TLam
(
App
(
Rec
(
App
(
Rec
(
CG_stack_body
(
Var
1
)
(
Var
3
)))
(
Alloc
(
Fold
(
InjL
Unit
)))))
newlock
).
Section
CG_Stack
.
...
...
@@ -82,7 +82,7 @@ Section CG_Stack.
={
E
}=>
j
⤇
fill
K
Unit
★
st
↦ₛ
FoldV
(
InjRV
(
PairV
w
v
)).
Proof
.
intros
HNE
.
iIntros
"[#Hspec [Hx Hj]]"
.
unfold
CG_push
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_load
_
_
j
(
K
++
[
StoreRCtx
(
LocV
_
);
FoldCtx
;
InjRCtx
;
PairRCtx
_
])
...
...
@@ -179,7 +179,7 @@ Section CG_Stack.
={
E
}=>
j
⤇
fill
K
(
InjR
(#
w
))
★
st
↦ₛ
v
.
Proof
.
intros
HNE
.
iIntros
"[#Hspec [Hx Hj]]"
.
unfold
CG_pop
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_load
_
_
j
(
K
++
[
CaseCtx
_
_
;
UnfoldCtx
])
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
...
...
@@ -193,16 +193,16 @@ Section CG_Stack.
rewrite
?
fill_app
.
simpl
.
iPvs
(
step_case_inr
_
_
j
K
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_snd
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
);
StoreRCtx
(
LocV
_
)])
_
_
_
_
iPvs
(
step_snd
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
);
StoreRCtx
(
LocV
_
)])
_
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
rewrite
?
fill_app
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
iPvs
(
step_store
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
)])
_
_
_
_
_
_
iPvs
(
step_store
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
)])
_
_
_
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
rewrite
?
fill_app
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
rewrite
?
fill_app
.
simpl
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_fst
_
_
j
(
K
++
[
InjRCtx
])
_
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
...
...
@@ -222,7 +222,7 @@ Section CG_Stack.
={
E
}=>
j
⤇
fill
K
(
InjL
Unit
)
★
st
↦ₛ
FoldV
(
InjLV
UnitV
).
Proof
.
iIntros
{
HNE
}
"[#Hspec [Hx Hj]]"
.
unfold
CG_pop
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_load
_
_
j
(
K
++
[
CaseCtx
_
_
;
UnfoldCtx
])
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
...
...
@@ -348,7 +348,7 @@ Section CG_Stack.
iPvs
(
steps_with_lock
_
_
j
K
_
_
_
_
v
UnitV
_
_
with
"[Hj Hx Hl]"
)
as
"Hj"
;
last
done
;
[
|
by
iFrame
"Hspec Hx Hl Hj"
].
iIntros
{
K
'
}
"[#Hspec [Hx Hj]]"
.
iPvs
(
step_
lam
_
_
j
K
'
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
'
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_load
_
_
j
K
'
_
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
...
...
@@ -364,10 +364,10 @@ Section CG_Stack.
(
*
Coarse
-
grained
iter
*
)
Lemma
CG_iter_folding
(
f
:
expr
)
:
CG_iter
f
=
Lam
(
Case
(
Unfold
(
Var
1
))
Rec
(
Case
(
Unfold
(
Var
1
))
Unit
(
App
(
Lam
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
App
f
.[
ren
(
+
3
)]
(
Fst
(
Var
0
)))
)
).
...
...
@@ -410,13 +410,13 @@ Section CG_Stack.
={
E
}=>
j
⤇
fill
K
(
App
(
Lam
(
Rec
(
App
((
CG_iter
(#
f
)).[
ren
(
+
2
)])
(
Snd
(
Pair
((#
w
).[
ren
(
+
2
)])
(#
v
).[
ren
(
+
2
)]))))
(
App
(#
f
)
(#
w
))).
Proof
.
iIntros
{
HNE
}
"[#Hspec Hj]"
.
unfold
CG_iter
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
rewrite
-
CG_iter_folding
.
Opaque
CG_iter
.
asimpl
.
iPvs
(
step_Fold
_
_
j
(
K
++
[
CaseCtx
_
_
])
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
...
...
@@ -425,7 +425,7 @@ Section CG_Stack.
rewrite
?
fill_app
.
asimpl
.
iPvs
(
step_case_inr
_
_
j
K
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
asimpl
.
iPvs
(
step_fst
_
_
j
(
K
++
[
AppRCtx
(
Lam
V
_
);
AppRCtx
f
])
_
_
_
_
iPvs
(
step_fst
_
_
j
(
K
++
[
AppRCtx
(
Rec
V
_
);
AppRCtx
f
])
_
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
rewrite
?
fill_app
/=
.
iFrame
"Hspec Hj"
;
trivial
.
...
...
@@ -443,7 +443,7 @@ Section CG_Stack.
={
E
}=>
j
⤇
fill
K
Unit
.
Proof
.
iIntros
{
HNE
}
"[#Hspec Hj]"
.
unfold
CG_iter
.
iPvs
(
step_
lam
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iPvs
(
step_
rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
rewrite
-
CG_iter_folding
.
Opaque
CG_iter
.
asimpl
.
iPvs
(
step_Fold
_
_
j
(
K
++
[
CaseCtx
_
_
])
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
...
...
@@ -509,7 +509,7 @@ Section CG_Stack.
Qed
.
Lemma
CG_stack_type
Γ
:
typed
Γ
(
CG_stack
)
typed
Γ
CG_stack
(
TForall
(
TProd
(
TProd
...
...
F_mu_ref_par/examples/stack/FG_stack.v
View file @
9a182da8
...
...
@@ -4,7 +4,7 @@ Definition FG_StackType τ :=
TRec
(
Tref
(
TSum
TUnit
(
TProd
τ
.[
ren
(
+
1
)]
(
TVar
0
)))).
Definition
FG_push
(
st
:
expr
)
:
expr
:=
Lam
(
App
(
Lam
Rec
(
App
(
Rec
(
*
try
push
*
)
(
If
(
CAS
(
st
.[
ren
(
+
4
)])
(
Var
1
)
(
Fold
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Var
1
)))))
...
...
@@ -16,7 +16,7 @@ Definition FG_push (st : expr) : expr :=
(
Load
st
.[
ren
(
+
2
)])
(
*
read
the
stack
pointer
*
)
).
Definition
FG_pushV
(
st
:
expr
)
:
val
:=
Lam
V
(
App
(
Lam
Rec
V
(
App
(
Rec
(
*
try
push
*
)
(
If
(
CAS
(
st
.[
ren
(
+
4
)])
(
Var
1
)
(
Fold
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Var
1
)))))
...
...
@@ -29,9 +29,9 @@ Definition FG_pushV (st : expr) : val :=
).
Definition
FG_pop
(
st
:
expr
)
:
expr
:=
Lam
(
App
(
Lam
Rec
(
App
(
Rec
(
App
(
Lam
(
Rec
(
Case
(
Var
1
)
(
InjL
Unit
)
...
...
@@ -55,11 +55,11 @@ Definition FG_pop (st : expr) : expr :=
(
Unfold
(
Load
st
.[
ren
(
+
2
)]))
).
Definition
FG_popV
(
st
:
expr
)
:
val
:=
Lam
V
Rec
V
(
App
(
Lam
(
Rec
(
App
(
Lam
(
Rec
(
Case
(
Var
1
)
(
InjL
Unit
)
...
...
@@ -84,36 +84,36 @@ Definition FG_popV (st : expr) : val :=
).
Definition
FG_iter
(
f
:
expr
)
:
expr
:=
Lam
Rec
(
Case
(
Load
(
Unfold
(
Var
1
)))
Unit
(
App
(
Lam
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
*
recursive_call
*
)
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
*
recursive_call
*
)
(
App
f
.[
ren
(
+
3
)]
(
Fst
(
Var
0
)))
)
).
Definition
FG_iterV
(
f
:
expr
)
:
val
:=
Lam
V
Rec
V
(
Case
(
Load
(
Unfold
(
Var
1
)))
Unit
(
App
(
Lam
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
*
recursive_call
*
)
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
*
recursive_call
*
)
(
App
f
.[
ren
(
+
3
)]
(
Fst
(
Var
0
)))
)
).
Definition
FG_read_iter
(
st
:
expr
)
:
expr
:=
Lam
(
App
(
FG_iter
(
Var
1
))
(
Load
st
.[
ren
(
+
2
)])).
Rec
(
App
(
FG_iter
(
Var
1
))
(
Load
st
.[
ren
(
+
2
)])).