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
Paolo G. Giarrusso
examples
Commits
76c60e8f
Commit
76c60e8f
authored
Dec 10, 2018
by
Amin Timany
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Simplify stack programs
parent
f72453a9
Changes
4
Hide whitespace changes
Inline
Side-by-side
Showing
4 changed files
with
429 additions
and
500 deletions
+429
-500
theories/logrel/F_mu_ref_conc/examples/lock.v
theories/logrel/F_mu_ref_conc/examples/lock.v
+51
-48
theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v
theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v
+151
-204
theories/logrel/F_mu_ref_conc/examples/stack/FG_stack.v
theories/logrel/F_mu_ref_conc/examples/stack/FG_stack.v
+198
-219
theories/logrel/F_mu_ref_conc/examples/stack/refinement.v
theories/logrel/F_mu_ref_conc/examples/stack/refinement.v
+29
-29
No files found.
theories/logrel/F_mu_ref_conc/examples/lock.v
View file @
76c60e8f
...
...
@@ -7,20 +7,26 @@ Definition newlock : expr := Alloc (#♭ false).
Definition
acquire
:
expr
:
=
Rec
(
If
(
CAS
(
Var
1
)
(#
♭
false
)
(#
♭
true
))
(
Unit
)
(
App
(
Var
0
)
(
Var
1
))).
(** [release = λ x. x <- false] *)
Definition
release
:
expr
:
=
Rec
(
Store
(
Var
1
)
(#
♭
false
)).
Definition
release
:
expr
:
=
Lam
(
Store
(
Var
0
)
(#
♭
false
)).
(** [with_lock e l = λ x. (acquire l) ;; e x ;; (release l)] *)
Definition
with_lock
(
e
:
expr
)
(
l
:
expr
)
:
expr
:
=
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
)])
Lam
(
Seq
(
App
acquire
l
.[
ren
(+
1
)])
(
LetIn
(
App
e
.[
ren
(+
1
)]
(
Var
0
))
(
Seq
(
App
release
l
.[
ren
(+
2
)])
(
Var
0
))
)
).
Definition
with_lockV
(
e
l
:
expr
)
:
val
:
=
RecV
(
App
(
Rec
(
App
(
Rec
(
App
(
Rec
(
Var
3
))
(
App
release
l
.[
ren
(+
6
)])))
(
App
e
.[
ren
(+
4
)]
(
Var
3
))))
(
App
acquire
l
.[
ren
(+
2
)])
LamV
(
Seq
(
App
acquire
l
.[
ren
(+
1
)])
(
LetIn
(
App
e
.[
ren
(+
1
)]
(
Var
0
))
(
Seq
(
App
release
l
.[
ren
(+
2
)])
(
Var
0
))
)
).
Lemma
with_lock_to_val
e
l
:
to_val
(
with_lock
e
l
)
=
Some
(
with_lockV
e
l
).
...
...
@@ -29,6 +35,7 @@ Proof. trivial. Qed.
Lemma
with_lock_of_val
e
l
:
of_val
(
with_lockV
e
l
)
=
with_lock
e
l
.
Proof
.
trivial
.
Qed
.
Global
Typeclasses
Opaque
with_lockV
.
Global
Opaque
with_lockV
.
Lemma
newlock_closed
f
:
newlock
.[
f
]
=
newlock
.
...
...
@@ -43,15 +50,10 @@ Lemma release_closed f : release.[f] = release.
Proof
.
by
asimpl
.
Qed
.
Hint
Rewrite
release_closed
:
autosubst
.
Lemma
with_lock_subst
(
e
l
:
expr
)
f
:
(
with_lock
e
l
).[
f
]
=
with_lock
e
.[
f
]
l
.[
f
].
Lemma
with_lock_subst
(
e
l
:
expr
)
f
:
(
with_lock
e
l
).[
f
]
=
with_lock
e
.[
f
]
l
.[
f
].
Proof
.
unfold
with_lock
;
asimpl
;
trivial
.
Qed
.
Lemma
with_lock_closed
e
l
:
(
∀
f
:
var
→
expr
,
e
.[
f
]
=
e
)
→
(
∀
f
:
var
→
expr
,
l
.[
f
]
=
l
)
→
∀
f
,
(
with_lock
e
l
).[
f
]
=
with_lock
e
l
.
Proof
.
asimpl
=>
H1
H2
f
.
unfold
with_lock
.
by
rewrite
?H1
?H2
.
Qed
.
Definition
LockType
:
=
Tref
TBool
.
Lemma
newlock_type
Γ
:
typed
Γ
newlock
LockType
.
...
...
@@ -68,18 +70,20 @@ Lemma with_lock_type e l Γ τ τ' :
typed
Γ
l
LockType
→
typed
Γ
(
with_lock
e
l
)
(
TArrow
τ
τ
'
).
Proof
.
intros
H1
H2
.
do
3
econstructor
;
eauto
.
-
repeat
(
econstructor
;
eauto
using
release_type
).
+
eapply
(
context_weakening
[
_;
_;
_;
_;
_;
_
])
;
eauto
.
+
eapply
(
context_weakening
[
_;
_;
_;
_
])
;
eauto
.
-
eapply
acquire_type
.
-
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
intros
??.
do
3
econstructor
;
eauto
using
acquire_type
.
-
eapply
(
context_weakening
[
_
])
;
eauto
.
-
econstructor
;
eauto
using
typed
.
eapply
(
context_weakening
[
_
])
;
eauto
.
-
econstructor
;
eauto
using
typed
.
econstructor
;
eauto
using
release_type
.
eapply
(
context_weakening
[
_;_
])
;
eauto
.
Qed
.
Section
proof
.
Context
`
{
cfgSG
Σ
}.
Context
`
{
heapIG
Σ
}.
Lemma
steps_newlock
E
ρ
j
K
:
nclose
specN
⊆
E
→
spec_ctx
ρ
∗
j
⤇
fill
K
newlock
...
...
@@ -89,6 +93,7 @@ Section proof.
by
iMod
(
step_alloc
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
.
Qed
.
Global
Typeclasses
Opaque
newlock
.
Global
Opaque
newlock
.
Lemma
steps_acquire
E
ρ
j
K
l
:
...
...
@@ -107,6 +112,7 @@ Section proof.
Unshelve
.
all
:
trivial
.
Qed
.
Global
Typeclasses
Opaque
acquire
.
Global
Opaque
acquire
.
Lemma
steps_release
E
ρ
j
K
l
b
:
...
...
@@ -115,48 +121,45 @@ Section proof.
⊢
|={
E
}=>
j
⤇
fill
K
Unit
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
iIntros
(
HNE
)
"[#Hspec [Hl Hj]]"
.
unfold
release
.
iMod
(
step_rec
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
;
try
done
.
iMod
(
step_store
_
_
j
K
_
_
_
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
{
by
iFrame
.
}
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
;
try
done
.
iMod
(
step_store
with
"[$Hj $Hl]"
)
as
"[Hj Hl]"
;
eauto
.
by
iIntros
"!> {$Hj $Hl}"
.
Unshelve
.
all
:
trivial
.
Qed
.
Global
Typeclasses
Opaque
release
.
Global
Opaque
release
.
Lemma
steps_with_lock
E
ρ
j
K
e
l
P
Q
v
w
:
nclose
specN
⊆
E
→
(
∀
f
,
e
.[
f
]
=
e
)
(* e is a closed term *)
→
(*
(∀ f, e.[f] = e) (* e is a closed term *) →
*)
(
∀
K'
,
spec_ctx
ρ
∗
P
∗
j
⤇
fill
K'
(
App
e
(
of_val
w
))
⊢
|={
E
}=>
j
⤇
fill
K'
(
of_val
v
)
∗
Q
)
→
spec_ctx
ρ
∗
P
∗
l
↦ₛ
(#
♭
v
false
)
∗
j
⤇
fill
K
(
App
(
with_lock
e
(
Loc
l
))
(
of_val
w
))
⊢
|={
E
}=>
j
⤇
fill
K
(
of_val
v
)
∗
Q
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
iIntros
(
HNE
H1
H2
)
"[#Hspec [HP [Hl Hj]]]"
.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
rewrite
H1
.
iMod
(
steps_acquire
_
_
j
((
AppRCtx
(
RecV
_
))
::
K
)
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
{
simpl
.
iFrame
"Hspec Hj Hl"
;
eauto
.
}
simpl
.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
rewrite
H1
.
iMod
(
H2
((
AppRCtx
(
RecV
_
))
::
K
)
with
"[Hj HP]"
)
as
"[Hj HQ]"
;
eauto
.
{
simpl
.
iFrame
"Hspec Hj HP"
;
eauto
.
}
iIntros
(
HNE
He
)
"[#Hspec [HP [Hl Hj]]]"
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
steps_acquire
_
_
j
(
SeqCtx
_
::
K
)
with
"[$Hj Hl]"
)
as
"[Hj Hl]"
;
auto
.
simpl
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
He
(
LetInCtx
_
::
K
)
with
"[$Hj HP]"
)
as
"[Hj HQ]"
;
eauto
.
simpl
.
iMod
(
step_re
c
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_
step_
pu
re
with
"[
$
Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
steps_release
_
_
j
(
(
AppRCtx
(
RecV
_
))
::
K
)
_
_
with
"[Hj Hl]"
)
iMod
(
steps_release
_
_
j
(
SeqCtx
_
::
K
)
_
_
with
"[
$
Hj
$
Hl]"
)
as
"[Hj Hl]"
;
eauto
.
{
simpl
.
by
iFrame
.
}
rewrite
?fill_app
/=.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iModIntro
;
by
iFrame
.
Unshelve
.
all
:
try
match
goal
with
|-
to_val
_
=
_
=>
auto
using
to_of_val
end
.
trivial
.
simpl
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iModIntro
;
by
iFrame
.
Qed
.
Global
Typeclasses
Opaque
with_lock
.
Global
Opaque
with_lock
.
End
proof
.
Global
Hint
Rewrite
newlock_closed
:
autosubst
.
Global
Hint
Rewrite
acquire_closed
:
autosubst
.
Global
Hint
Rewrite
release_closed
:
autosubst
.
Global
Hint
Rewrite
with_lock_subst
:
autosubst
.
theories/logrel/F_mu_ref_conc/examples/stack/CG_stack.v
View file @
76c60e8f
...
...
@@ -7,33 +7,34 @@ Definition CG_StackType τ :=
(* Coarse-grained push *)
Definition
CG_push
(
st
:
expr
)
:
expr
:
=
Rec
(
Store
(
st
.[
ren
(+
2
)])
(
Fold
(
InjR
(
Pair
(
Var
1
)
(
Load
st
.[
ren
(+
2
)]))))).
Lam
(
Store
(
st
.[
ren
(+
1
)])
(
Fold
(
InjR
(
Pair
(
Var
0
)
(
Load
st
.[
ren
(+
1
)]))))).
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
:
=
Rec
(
Case
(
Unfold
(
Load
st
.[
ren
(+
2
)]))
Lam
(
Case
(
Unfold
(
Load
st
.[
ren
(+
1
)]))
(
InjL
Unit
)
(
App
(
Rec
(
InjR
(
Fst
(
Var
2
))))
(
Store
st
.[
ren
(+
3
)]
(
Snd
(
Var
0
)))
Seq
(
Store
st
.[
ren
(+
2
)]
(
Snd
(
Var
0
)))
(
InjR
(
Fst
(
Var
0
)))
)
).
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
(
Rec
(
Load
st
.[
ren
(+
2
)]))
l
.
Definition
CG_snapV
(
st
l
:
expr
)
:
val
:
=
with_lockV
(
Rec
(
Load
st
.[
ren
(+
2
)]))
l
.
Definition
CG_snap
(
st
l
:
expr
)
:
=
with_lock
(
Lam
(
Load
st
.[
ren
(+
1
)]))
l
.
Definition
CG_snapV
(
st
l
:
expr
)
:
val
:
=
with_lockV
(
Lam
(
Load
st
.[
ren
(+
1
)]))
l
.
Definition
CG_iter
(
f
:
expr
)
:
expr
:
=
Rec
(
Case
(
Unfold
(
Var
1
))
Unit
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
Seq
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
(
Var
1
)
(
Snd
(
Var
0
)))
)
).
...
...
@@ -41,20 +42,26 @@ Definition CG_iterV (f : expr) : val :=
RecV
(
Case
(
Unfold
(
Var
1
))
Unit
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
Seq
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
(
Var
1
)
(
Snd
(
Var
0
)))
)
).
Definition
CG_snap_iter
(
st
l
:
expr
)
:
expr
:
=
Rec
(
App
(
CG_iter
(
Var
1
))
(
App
(
CG_snap
st
.[
ren
(+
2
)]
l
.[
ren
(+
2
)])
Unit
)).
Lam
(
App
(
CG_iter
(
Var
0
))
(
App
(
CG_snap
st
.[
ren
(+
1
)]
l
.[
ren
(+
1
)])
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
(
Rec
(
App
(
Rec
(
CG_stack_body
(
Var
1
)
(
Var
3
)))
(
Alloc
(
Fold
(
InjL
Unit
)))))
newlock
).
TLam
(
LetIn
newlock
(
LetIn
(
Alloc
(
Fold
(
InjL
Unit
)))
(
CG_stack_body
(
Var
0
)
(
Var
1
)))).
Section
CG_Stack
.
Context
`
{
heapIG
Σ
,
cfgSG
Σ
}.
...
...
@@ -64,37 +71,33 @@ Section CG_Stack.
typed
Γ
(
CG_push
st
)
(
TArrow
τ
TUnit
).
Proof
.
intros
H1
.
repeat
econstructor
.
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
eapply
(
context_weakening
[
_
])
;
eauto
.
repeat
constructor
;
asimpl
;
trivial
.
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
eapply
(
context_weakening
[
_
])
;
eauto
.
Qed
.
Lemma
CG_push_closed
(
st
:
expr
)
:
(
∀
f
,
st
.[
f
]
=
st
)
→
∀
f
,
(
CG_push
st
).[
f
]
=
CG_push
st
.
Proof
.
intros
Hst
f
.
unfold
CG_push
.
asimpl
.
rewrite
?Hst
;
trivial
.
Qed
.
Lemma
CG_push_subst
(
st
:
expr
)
f
:
(
CG_push
st
).[
f
]
=
CG_push
st
.[
f
].
Proof
.
unfold
CG_push
;
asimpl
;
trivial
.
Qed
.
Global
Hint
Rewrite
CG_push_subst
:
autosubst
.
Lemma
steps_CG_push
E
ρ
j
K
st
v
w
:
nclose
specN
⊆
E
→
spec_ctx
ρ
∗
st
↦ₛ
v
∗
j
⤇
fill
K
(
App
(
CG_push
(
Loc
st
))
(
of_val
w
))
⊢
|={
E
}=>
j
⤇
fill
K
Unit
∗
st
↦ₛ
FoldV
(
InjRV
(
PairV
w
v
)).
Proof
.
intros
HNE
.
iIntros
"[#Hspec [Hx Hj]]"
.
unfold
CG_push
.
iMod
(
step_re
c
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_
step_
pu
re
with
"[
$
Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
step_load
_
_
j
(
PairRCtx
_
::
InjRCtx
::
FoldCtx
::
StoreRCtx
(
LocV
_
)
::
K
)
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
simpl
.
iMod
(
step_store
_
_
j
K
_
_
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
{
by
iFrame
.
}
with
"[
$
Hj
$
Hx]"
)
as
"[Hj Hx]"
;
eauto
.
simpl
.
iMod
(
step_store
_
_
j
K
with
"[
$
Hj
$
Hx]"
)
as
"[Hj Hx]"
;
eauto
.
{
rewrite
/=
!
to_of_val
//
.
}
iModIntro
.
by
iFrame
.
Unshelve
.
all
:
try
match
goal
with
|-
to_val
_
=
_
=>
auto
using
to_of_val
end
.
simpl
;
by
rewrite
?to_of_val
.
Qed
.
Global
Typeclasses
Opaque
CG_push
.
Global
Opaque
CG_push
.
Lemma
CG_locked_push_to_val
st
l
:
...
...
@@ -106,7 +109,6 @@ Section CG_Stack.
Proof
.
trivial
.
Qed
.
Global
Opaque
CG_locked_pushV
.
Lemma
CG_locked_push_type
st
l
Γ
τ
:
typed
Γ
st
(
Tref
(
CG_StackType
τ
))
→
typed
Γ
l
LockType
→
...
...
@@ -116,17 +118,11 @@ Section CG_Stack.
eapply
with_lock_type
;
auto
using
CG_push_type
.
Qed
.
Lemma
CG_locked_push_closed
(
st
l
:
expr
)
:
(
∀
f
,
st
.[
f
]
=
st
)
→
(
∀
f
,
l
.[
f
]
=
l
)
→
∀
f
,
(
CG_locked_push
st
l
).[
f
]
=
CG_locked_push
st
l
.
Proof
.
intros
H1
H2
f
.
asimpl
.
unfold
CG_locked_push
.
rewrite
with_lock_closed
;
trivial
.
apply
CG_push_closed
;
trivial
.
Qed
.
Lemma
CG_locked_push_subst
(
st
l
:
expr
)
f
:
(
CG_locked_push
st
l
).[
f
]
=
CG_locked_push
st
.[
f
]
l
.[
f
].
Proof
.
by
rewrite
with_lock_subst
CG_push_subst
.
Qed
.
Proof
.
by
rewrite
/
CG_locked_push
;
asimpl
.
Qed
.
Hint
Rewrite
CG_locked_push_subst
:
autosubst
.
Lemma
steps_CG_locked_push
E
ρ
j
K
st
w
v
l
:
nclose
specN
⊆
E
→
...
...
@@ -135,14 +131,13 @@ Section CG_Stack.
⊢
|={
E
}=>
j
⤇
fill
K
Unit
∗
st
↦ₛ
FoldV
(
InjRV
(
PairV
w
v
))
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
intros
HNE
.
iIntros
"[#Hspec [Hx [Hl Hj]]]"
.
unfold
CG_locked_push
.
iMod
(
steps_with_lock
_
_
j
K
_
_
_
_
UnitV
_
_
_
with
"[Hj Hx Hl]"
)
as
"Hj"
;
last
done
.
-
iIntros
(
K'
)
"[#Hspec Hxj]"
.
iApply
steps_CG_push
;
first
done
.
iFrame
.
trivial
.
-
iFrame
"Hspec Hj Hx"
;
trivial
.
Unshelve
.
all
:
trivial
.
iMod
(
steps_with_lock
_
_
_
_
_
_
(
st
↦ₛ
v
)%
I
_
UnitV
with
"[$Hj $Hx $Hl]"
)
as
"Hj"
;
auto
.
iIntros
(
K'
)
"[#Hspec Hxj]"
.
iApply
steps_CG_push
;
first
done
.
by
iFrame
.
Qed
.
Global
Typeclasses
Opaque
CG_locked_push
.
Global
Opaque
CG_locked_push
.
(* Coarse-grained pop *)
...
...
@@ -152,22 +147,22 @@ Section CG_Stack.
Proof
.
intros
H1
.
econstructor
.
eapply
(
Case_typed
_
_
_
_
TUnit
)
;
[|
repeat
constructor
|
repeat
econstructor
;
eapply
(
context_weakening
[
_;
_;
_
])
;
eauto
].
replace
(
TSum
TUnit
(
TProd
τ
(
CG_StackType
τ
)))
with
((
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
))).[(
CG_StackType
τ
)/])
by
(
by
asimpl
).
repeat
econstructor
.
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
eapply
(
Case_typed
_
_
_
_
TUnit
)
;
eauto
using
typed
.
-
replace
(
TSum
TUnit
(
TProd
τ
(
CG_StackType
τ
)))
with
((
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
))).[(
CG_StackType
τ
)/])
by
(
by
asimpl
).
repeat
econstructor
.
eapply
(
context_weakening
[
_
])
;
eauto
.
-
econstructor
;
eauto
using
typed
.
econstructor
;
eauto
using
typed
.
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
Qed
.
Lemma
CG_pop_
closed
(
st
:
expr
)
:
(
∀
f
,
st
.[
f
]
=
st
)
→
∀
f
,
(
CG_pop
st
).[
f
]
=
CG_pop
st
.
Proof
.
intros
Hst
f
.
unfold
CG_pop
.
asimpl
.
rewrite
?Hst
;
trivial
.
Qed
.
Lemma
CG_pop_
subst
(
st
:
expr
)
f
:
(
CG_pop
st
).[
f
]
=
CG_pop
st
.
[
f
].
Proof
.
by
rewrite
/
CG_pop
;
asimpl
.
Qed
.
Lemma
CG_pop_subst
(
st
:
expr
)
f
:
(
CG_pop
st
).[
f
]
=
CG_pop
st
.[
f
].
Proof
.
unfold
CG_pop
;
asimpl
;
trivial
.
Qed
.
Hint
Rewrite
CG_pop_subst
:
autosubst
.
Lemma
steps_CG_pop_suc
E
ρ
j
K
st
v
w
:
nclose
specN
⊆
E
→
...
...
@@ -176,35 +171,25 @@ Section CG_Stack.
⊢
|={
E
}=>
j
⤇
fill
K
(
InjR
(
of_val
w
))
∗
st
↦ₛ
v
.
Proof
.
intros
HNE
.
iIntros
"[#Hspec [Hx Hj]]"
.
unfold
CG_pop
.
iMod
(
step_re
c
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_
step_
pu
re
with
"[
$
Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
step_load
_
_
j
(
UnfoldCtx
::
CaseCtx
_
_
::
K
)
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
rewrite
?fill_app
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
rewrite
?fill_app
.
simpl
.
iMod
(
step_Fold
_
_
j
(
CaseCtx
_
_
::
K
)
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
step_load
_
_
_
(
UnfoldCtx
::
CaseCtx
_
_
::
K
)
with
"[$Hj $Hx]"
)
as
"[Hj Hx]"
;
eauto
.
simpl
.
iMod
(
step_case_inr
_
_
j
K
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
step_snd
_
_
j
(
StoreRCtx
(
LocV
_
)
::
AppRCtx
(
RecV
_
)
::
K
)
_
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_step_pure
_
_
_
(
CaseCtx
_
_
::
K
)
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
simpl
.
iMod
(
step_store
_
_
j
(
AppRCtx
(
RecV
_
)
::
K
)
_
_
_
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
rewrite
?fill_app
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
rewrite
?fill_app
.
simpl
.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
step_
fst
_
_
j
(
InjRCtx
::
K
)
_
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_
step_
pure
_
_
_
(
StoreRCtx
(
LocV
_
)
::
SeqCtx
_
::
K
)
with
"[
$
Hj]"
)
as
"Hj"
;
eauto
.
simpl
.
iModIntro
.
iFrame
"Hj Hx"
;
trivial
.
Unshelve
.
all
:
try
match
goal
with
|-
to_val
_
=
_
=>
simpl
;
by
rewrite
?to_of_val
end
.
all
:
trivial
.
iMod
(
step_store
_
_
j
(
SeqCtx
_
::
K
)
with
"[$Hj $Hx]"
)
as
"[Hj Hx]"
;
eauto
using
to_of_val
.
simpl
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_step_pure
_
_
_
(
InjRCtx
::
K
)
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
simpl
.
by
iFrame
"Hj Hx"
.
Qed
.
Lemma
steps_CG_pop_fail
E
ρ
j
K
st
:
...
...
@@ -214,21 +199,17 @@ Section CG_Stack.
⊢
|={
E
}=>
j
⤇
fill
K
(
InjL
Unit
)
∗
st
↦ₛ
FoldV
(
InjLV
UnitV
).
Proof
.
iIntros
(
HNE
)
"[#Hspec [Hx Hj]]"
.
unfold
CG_pop
.
iMod
(
step_re
c
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_
step_
pu
re
with
"[
$
Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iMod
(
step_load
_
_
j
(
UnfoldCtx
::
CaseCtx
_
_
::
K
)
_
_
_
with
"[Hj Hx]"
)
as
"[Hj Hx]"
;
eauto
.
simpl
.
iFrame
"Hspec Hj"
;
trivial
.
simpl
.
iMod
(
step_Fold
_
_
j
(
CaseCtx
_
_
::
K
)
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
step_case_inl
_
_
j
K
_
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iModIntro
.
iFrame
"Hj Hx"
;
trivial
.
Unshelve
.
all
:
try
match
goal
with
|-
to_val
_
=
_
=>
simpl
;
by
rewrite
?to_of_val
end
.
all
:
trivial
.
_
_
_
with
"[$Hj $Hx]"
)
as
"[Hj Hx]"
;
eauto
.
iMod
(
do_step_pure
_
_
_
(
CaseCtx
_
_
::
K
)
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
simpl
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
by
iFrame
"Hj Hx"
;
trivial
.
Qed
.
Global
Typeclasses
Opaque
CG_pop
.
Global
Opaque
CG_pop
.
Lemma
CG_locked_pop_to_val
st
l
:
...
...
@@ -250,17 +231,11 @@ Section CG_Stack.
eapply
with_lock_type
;
auto
using
CG_pop_type
.
Qed
.
Lemma
CG_locked_pop_closed
(
st
l
:
expr
)
:
(
∀
f
,
st
.[
f
]
=
st
)
→
(
∀
f
,
l
.[
f
]
=
l
)
→
∀
f
,
(
CG_locked_pop
st
l
).[
f
]
=
CG_locked_pop
st
l
.
Proof
.
intros
H1
H2
f
.
asimpl
.
unfold
CG_locked_pop
.
rewrite
with_lock_closed
;
trivial
.
apply
CG_pop_closed
;
trivial
.
Qed
.
Lemma
CG_locked_pop_subst
(
st
l
:
expr
)
f
:
(
CG_locked_pop
st
l
).[
f
]
=
CG_locked_pop
st
.[
f
]
l
.[
f
].
Proof
.
by
rewrite
with_lock_subst
CG_pop_subst
.
Qed
.
(
CG_locked_pop
st
l
).[
f
]
=
CG_locked_pop
st
.[
f
]
l
.[
f
].
Proof
.
by
rewrite
/
CG_locked_pop
;
asimpl
.
Qed
.
Hint
Rewrite
CG_locked_pop_subst
:
autosubst
.
Lemma
steps_CG_locked_pop_suc
E
ρ
j
K
st
v
w
l
:
nclose
specN
⊆
E
→
...
...
@@ -269,12 +244,11 @@ Section CG_Stack.
⊢
|={
E
}=>
j
⤇
fill
K
(
InjR
(
of_val
w
))
∗
st
↦ₛ
v
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
iIntros
(
HNE
)
"[#Hspec [Hx [Hl Hj]]]"
.
unfold
CG_locked_pop
.
iMod
(
steps_with_lock
_
_
j
K
_
_
_
_
(
InjRV
w
)
UnitV
_
_
with
"[Hj Hx Hl]"
)
as
"Hj"
;
last
done
.
-
iIntros
(
K'
)
"[#Hspec Hxj]"
.
iApply
steps_CG_pop_suc
;
first
done
.
by
iFrame
.
-
iFrame
"Hspec Hj Hx"
;
trivial
.
Unshelve
.
all
:
trivial
.
iMod
(
steps_with_lock
_
_
_
_
_
_
(
st
↦ₛ
FoldV
(
InjRV
(
PairV
w
v
)))%
I
_
(
InjRV
w
)
UnitV
with
"[$Hj $Hx $Hl]"
)
as
"Hj"
;
eauto
.
iIntros
(
K'
)
"[#Hspec Hxj]"
.
iApply
steps_CG_pop_suc
;
eauto
.
Qed
.
Lemma
steps_CG_locked_pop_fail
E
ρ
j
K
st
l
:
...
...
@@ -284,14 +258,14 @@ Section CG_Stack.
⊢
|={
E
}=>
j
⤇
fill
K
(
InjL
Unit
)
∗
st
↦ₛ
FoldV
(
InjLV
UnitV
)
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
iIntros
(
HNE
)
"[#Hspec [Hx [Hl Hj]]]"
.
unfold
CG_locked_pop
.
iMod
(
steps_with_lock
_
_
j
K
_
_
_
_
(
InjLV
UnitV
)
UnitV
_
_
with
"[Hj Hx Hl]"
)
as
"Hj"
;
last
done
.
-
iIntros
(
K'
)
"[#Hspec Hxj] /="
.
iApply
steps_CG_pop_fail
;
first
done
.
by
iFrame
.
-
iFrame
"Hspec Hj Hx"
;
trivial
.
Unshelve
.
all
:
trivial
.
iMod
(
steps_with_lock
_
_
_
_
_
_
(
st
↦ₛ
FoldV
(
InjLV
UnitV
))%
I
_
(
InjLV
UnitV
)
UnitV
with
"[$Hj $Hx $Hl]"
)
as
"Hj"
;
eauto
.
iIntros
(
K'
)
"[#Hspec Hxj] /="
.
iApply
steps_CG_pop_fail
;
eauto
.
Qed
.
Global
Typeclasses
Opaque
CG_locked_pop
.
Global
Opaque
CG_locked_pop
.
Lemma
CG_snap_to_val
st
l
:
to_val
(
CG_snap
st
l
)
=
Some
(
CG_snapV
st
l
).
...
...
@@ -300,6 +274,7 @@ Section CG_Stack.
Lemma
CG_snap_of_val
st
l
:
of_val
(
CG_snapV
st
l
)
=
CG_snap
st
l
.
Proof
.
trivial
.
Qed
.
Global
Typeclasses
Opaque
CG_snapV
.
Global
Opaque
CG_snapV
.
Lemma
CG_snap_type
st
l
Γ
τ
:
...
...
@@ -309,21 +284,14 @@ Section CG_Stack.
Proof
.
intros
H1
H2
.
repeat
econstructor
.
eapply
with_lock_type
;
trivial
.
do
2
constructor
.
eapply
(
context_weakening
[
_;
_
])
;
eauto
.