Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
E
examples
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
0
Issues
0
List
Boards
Labels
Service Desk
Milestones
Merge Requests
0
Merge Requests
0
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Snippets
Snippets
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
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
Expand all
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).
...
@@ -7,20 +7,26 @@ Definition newlock : expr := Alloc (#♭ false).
Definition
acquire
:
expr
:
=
Definition
acquire
:
expr
:
=
Rec
(
If
(
CAS
(
Var
1
)
(#
♭
false
)
(#
♭
true
))
(
Unit
)
(
App
(
Var
0
)
(
Var
1
))).
Rec
(
If
(
CAS
(
Var
1
)
(#
♭
false
)
(#
♭
true
))
(
Unit
)
(
App
(
Var
0
)
(
Var
1
))).
(** [release = λ x. x <- false] *)
(** [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)] *)
(** [with_lock e l = λ x. (acquire l) ;; e x ;; (release l)] *)
Definition
with_lock
(
e
:
expr
)
(
l
:
expr
)
:
expr
:
=
Definition
with_lock
(
e
:
expr
)
(
l
:
expr
)
:
expr
:
=
Rec
Lam
(
App
(
Rec
(
App
(
Rec
(
App
(
Rec
(
Var
3
))
(
App
release
l
.[
ren
(+
6
)])))
(
Seq
(
App
e
.[
ren
(+
4
)]
(
Var
3
))))
(
App
acquire
l
.[
ren
(+
1
)])
(
App
acquire
l
.[
ren
(+
2
)])
(
LetIn
(
App
e
.[
ren
(+
1
)]
(
Var
0
))
(
Seq
(
App
release
l
.[
ren
(+
2
)])
(
Var
0
))
)
).
).
Definition
with_lockV
(
e
l
:
expr
)
:
val
:
=
Definition
with_lockV
(
e
l
:
expr
)
:
val
:
=
RecV
LamV
(
App
(
Rec
(
App
(
Rec
(
App
(
Rec
(
Var
3
))
(
App
release
l
.[
ren
(+
6
)])))
(
Seq
(
App
e
.[
ren
(+
4
)]
(
Var
3
))))
(
App
acquire
l
.[
ren
(+
1
)])
(
App
acquire
l
.[
ren
(+
2
)])
(
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
).
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.
...
@@ -29,6 +35,7 @@ Proof. trivial. Qed.
Lemma
with_lock_of_val
e
l
:
of_val
(
with_lockV
e
l
)
=
with_lock
e
l
.
Lemma
with_lock_of_val
e
l
:
of_val
(
with_lockV
e
l
)
=
with_lock
e
l
.
Proof
.
trivial
.
Qed
.
Proof
.
trivial
.
Qed
.
Global
Typeclasses
Opaque
with_lockV
.
Global
Opaque
with_lockV
.
Global
Opaque
with_lockV
.
Lemma
newlock_closed
f
:
newlock
.[
f
]
=
newlock
.
Lemma
newlock_closed
f
:
newlock
.[
f
]
=
newlock
.
...
@@ -43,15 +50,10 @@ Lemma release_closed f : release.[f] = release.
...
@@ -43,15 +50,10 @@ Lemma release_closed f : release.[f] = release.
Proof
.
by
asimpl
.
Qed
.
Proof
.
by
asimpl
.
Qed
.
Hint
Rewrite
release_closed
:
autosubst
.
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
.
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
.
Definition
LockType
:
=
Tref
TBool
.
Lemma
newlock_type
Γ
:
typed
Γ
newlock
LockType
.
Lemma
newlock_type
Γ
:
typed
Γ
newlock
LockType
.
...
@@ -68,18 +70,20 @@ Lemma with_lock_type e l Γ τ τ' :
...
@@ -68,18 +70,20 @@ Lemma with_lock_type e l Γ τ τ' :
typed
Γ
l
LockType
→
typed
Γ
l
LockType
→
typed
Γ
(
with_lock
e
l
)
(
TArrow
τ
τ
'
).
typed
Γ
(
with_lock
e
l
)
(
TArrow
τ
τ
'
).
Proof
.
Proof
.
intros
H1
H2
.
do
3
econstructor
;
eauto
.
intros
??.
-
repeat
(
econstructor
;
eauto
using
release_type
).
do
3
econstructor
;
eauto
using
acquire_type
.
+
eapply
(
context_weakening
[
_;
_;
_;
_;
_;
_
])
;
eauto
.
-
eapply
(
context_weakening
[
_
])
;
eauto
.
+
eapply
(
context_weakening
[
_;
_;
_;
_
])
;
eauto
.
-
econstructor
;
eauto
using
typed
.
-
eapply
acquire_type
.
eapply
(
context_weakening
[
_
])
;
eauto
.
-
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
-
econstructor
;
eauto
using
typed
.
econstructor
;
eauto
using
release_type
.
eapply
(
context_weakening
[
_;_
])
;
eauto
.
Qed
.
Qed
.
Section
proof
.
Section
proof
.
Context
`
{
cfgSG
Σ
}.
Context
`
{
cfgSG
Σ
}.
Context
`
{
heapIG
Σ
}.
Context
`
{
heapIG
Σ
}.
Lemma
steps_newlock
E
ρ
j
K
:
Lemma
steps_newlock
E
ρ
j
K
:
nclose
specN
⊆
E
→
nclose
specN
⊆
E
→
spec_ctx
ρ
∗
j
⤇
fill
K
newlock
spec_ctx
ρ
∗
j
⤇
fill
K
newlock
...
@@ -89,6 +93,7 @@ Section proof.
...
@@ -89,6 +93,7 @@ Section proof.
by
iMod
(
step_alloc
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
.
by
iMod
(
step_alloc
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
.
Qed
.
Qed
.
Global
Typeclasses
Opaque
newlock
.
Global
Opaque
newlock
.
Global
Opaque
newlock
.
Lemma
steps_acquire
E
ρ
j
K
l
:
Lemma
steps_acquire
E
ρ
j
K
l
:
...
@@ -107,6 +112,7 @@ Section proof.
...
@@ -107,6 +112,7 @@ Section proof.
Unshelve
.
all
:
trivial
.
Unshelve
.
all
:
trivial
.
Qed
.
Qed
.
Global
Typeclasses
Opaque
acquire
.
Global
Opaque
acquire
.
Global
Opaque
acquire
.
Lemma
steps_release
E
ρ
j
K
l
b
:
Lemma
steps_release
E
ρ
j
K
l
b
:
...
@@ -115,48 +121,45 @@ Section proof.
...
@@ -115,48 +121,45 @@ Section proof.
⊢
|={
E
}=>
j
⤇
fill
K
Unit
∗
l
↦ₛ
(#
♭
v
false
).
⊢
|={
E
}=>
j
⤇
fill
K
Unit
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
Proof
.
iIntros
(
HNE
)
"[#Hspec [Hl Hj]]"
.
unfold
release
.
iIntros
(
HNE
)
"[#Hspec [Hl Hj]]"
.
unfold
release
.
iMod
(
step_rec
_
_
j
K
with
"[Hj]"
)
as
"Hj"
;
eauto
;
try
done
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
;
try
done
.
iMod
(
step_store
_
_
j
K
_
_
_
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
iMod
(
step_store
with
"[$Hj $Hl]"
)
as
"[Hj Hl]"
;
eauto
.
{
by
iFrame
.
}
by
iIntros
"!> {$Hj $Hl}"
.
by
iIntros
"!> {$Hj $Hl}"
.
Unshelve
.
all
:
trivial
.
Qed
.
Qed
.
Global
Typeclasses
Opaque
release
.
Global
Opaque
release
.
Global
Opaque
release
.
Lemma
steps_with_lock
E
ρ
j
K
e
l
P
Q
v
w
:
Lemma
steps_with_lock
E
ρ
j
K
e
l
P
Q
v
w
:
nclose
specN
⊆
E
→
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
))
(
∀
K'
,
spec_ctx
ρ
∗
P
∗
j
⤇
fill
K'
(
App
e
(
of_val
w
))
⊢
|={
E
}=>
j
⤇
fill
K'
(
of_val
v
)
∗
Q
)
→
⊢
|={
E
}=>
j
⤇
fill
K'
(
of_val
v
)
∗
Q
)
→
spec_ctx
ρ
∗
P
∗
l
↦ₛ
(#
♭
v
false
)
spec_ctx
ρ
∗
P
∗
l
↦ₛ
(#
♭
v
false
)
∗
j
⤇
fill
K
(
App
(
with_lock
e
(
Loc
l
))
(
of_val
w
))
∗
j
⤇
fill
K
(
App
(
with_lock
e
(
Loc
l
))
(
of_val
w
))
⊢
|={
E
}=>
j
⤇
fill
K
(
of_val
v
)
∗
Q
∗
l
↦ₛ
(#
♭
v
false
).
⊢
|={
E
}=>
j
⤇
fill
K
(
of_val
v
)
∗
Q
∗
l
↦ₛ
(#
♭
v
false
).
Proof
.
Proof
.
iIntros
(
HNE
H1
H2
)
"[#Hspec [HP [Hl Hj]]]"
.
iIntros
(
HNE
He
)
"[#Hspec [HP [Hl Hj]]]"
.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
rewrite
H1
.
iAsimpl
.
iMod
(
steps_acquire
_
_
j
((
AppRCtx
(
RecV
_
))
::
K
)
iMod
(
steps_acquire
_
_
j
(
SeqCtx
_
::
K
)
with
"[$Hj Hl]"
)
as
"[Hj Hl]"
;
_
_
with
"[Hj Hl]"
)
as
"[Hj Hl]"
;
eauto
.
auto
.
simpl
.
{
simpl
.
iFrame
"Hspec Hj Hl"
;
eauto
.
}
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
simpl
.
iMod
(
He
(
LetInCtx
_
::
K
)
with
"[$Hj HP]"
)
as
"[Hj HQ]"
;
eauto
.
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
.
}
simpl
.
simpl
.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[
Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
do_step_pure
with
"[$
Hj]"
)
as
"Hj"
;
eauto
.
iAsimpl
.
iAsimpl
.
iMod
(
steps_release
_
_
j
(
(
AppRCtx
(
RecV
_
))
::
K
)
_
_
with
"[Hj
Hl]"
)
iMod
(
steps_release
_
_
j
(
SeqCtx
_
::
K
)
_
_
with
"[$Hj $
Hl]"
)
as
"[Hj Hl]"
;
eauto
.
as
"[Hj Hl]"
;
eauto
.
{
simpl
.
by
iFrame
.
}
simpl
.
rewrite
?fill_app
/=.
iMod
(
do_step_pure
with
"[$Hj]"
)
as
"Hj"
;
eauto
.
iMod
(
step_rec
_
_
j
K
_
_
_
_
with
"[Hj]"
)
as
"Hj"
;
eauto
.
iModIntro
;
by
iFrame
.
iAsimpl
.
iModIntro
;
by
iFrame
.
Unshelve
.
all
:
try
match
goal
with
|-
to_val
_
=
_
=>
auto
using
to_of_val
end
.
trivial
.
Qed
.
Qed
.
Global
Typeclasses
Opaque
with_lock
.
Global
Opaque
with_lock
.
Global
Opaque
with_lock
.
End
proof
.
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
This diff is collapsed.
Click to expand it.
theories/logrel/F_mu_ref_conc/examples/stack/FG_stack.v
View file @
76c60e8f
From
iris_examples
.
logrel
.
F_mu_ref_conc
Require
Import
typing
.
From
iris_examples
.
logrel
.
F_mu_ref_conc
Require
Import
typing
.
Definition
FG_StackType
τ
:
=
Definition
FG_StackType
τ
:
=
TRec
(
Tref
(
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
)))).
TRec
(
Tref
(
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
)))).
...
@@ -7,103 +7,99 @@ Definition FG_Stack_Ref_Type τ :=
...
@@ -7,103 +7,99 @@ Definition FG_Stack_Ref_Type τ :=
Tref
(
TSum
TUnit
(
TProd
τ
(
FG_StackType
τ
))).
Tref
(
TSum
TUnit
(
TProd
τ
(
FG_StackType
τ
))).
Definition
FG_push
(
st
:
expr
)
:
expr
:
=
Definition
FG_push
(
st
:
expr
)
:
expr
:
=
Rec
(
App
(
Rec
Rec
(* try push *)
(
LetIn
(
If
(
CAS
(
st
.[
ren
(+
4
)])
(
Var
1
)
(
Load
st
.[
ren
(+
2
)])
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Fold
(
Var
1
)))))
(* try push *)
)
(
If
(
CAS
(
st
.[
ren
(+
3
)])
(
Var
0
)
Unit
(* push succeeds we return unit *)
(
Alloc
(
InjR
(
Pair
(
Var
2
)
(
Fold
(
Var
0
)))))
(
App
(
Var
2
)
(
Var
3
))
(* push fails, we try again *)
)
)
)
(
Load
st
.[
ren
(+
2
)])
(* read the stack pointer *)
Unit
(* push succeeds we return unit *)
).
(
App
(
Var
1
)
(
Var
2
))
(* push fails, we try again *)
)
).
Definition
FG_pushV
(
st
:
expr
)
:
val
:
=
Definition
FG_pushV
(
st
:
expr
)
:
val
:
=
RecV
(
App
(
Rec
RecV
(* try push *)
(
LetIn
(
If
(
CAS
(
st
.[
ren
(+
4
)])
(
Var
1
)
(
Load
st
.[
ren
(+
2
)])
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Fold
(
Var
1
)))))
(* try push *)
)
(
If
(
CAS
(
st
.[
ren
(+
3
)])
(
Var
0
)
Unit
(* push succeeds we return unit *)
(
Alloc
(
InjR
(
Pair
(
Var
2
)
(
Fold
(
Var
0
)))))
(
App
(
Var
2
)
(
Var
3
))
(* push fails, we try again *)
)
)
)
(
Load
st
.[
ren
(+
2
)])
(* read the stack pointer *)
Unit
(* push succeeds we return unit *)
).
(
App
(
Var
1
)
(
Var
2
))
(* push fails, we try again *)
)
).
Definition
FG_pop
(
st
:
expr
)
:
expr
:
=
Definition
FG_pop
(
st
:
expr
)
:
expr
:
=
Rec
(
App
(
Rec
Rec
(
App
(
LetIn
(
Rec
(
Load
st
.[
ren
(+
2
)])
(
(
LetIn
Case
(
Var
1
)
(
Load
(
Var
0
))
(
InjL
Unit
)
(
Case
(
(* try popping *)
(
Var
0
)
If
(
InjL
Unit
)
(
CAS
(
(* try popping *)
(
st
.[
ren
(+
7
)])
If
(
Var
4
)
(
CAS
(
Unfold
(
Snd
(
Var
0
)))
(
st
.[
ren
(+
5
)])
)
(
Var
2
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
Unfold
(
Snd
(
Var
0
)))
(
App
(
Var
5
)
(
Var
6
))
(* pop fails, we retry*)
)
)
)
(
(
Load
(
Var
1
))
)
)
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
)
(
App
(
Var
3
)
(
Var
4
))
(* pop fails, we retry*)
(
Load
st
.[
ren
(+
2
)])
)
).
)
)
).
Definition
FG_popV
(
st
:
expr
)
:
val
:
=
Definition
FG_popV
(
st
:
expr
)
:
val
:
=
RecV
RecV
(
App
(
LetIn
(
Rec
(
Load
st
.[
ren
(+
2
)])
(
App
(
LetIn
(
Rec
(
Load
(
Var
0
))
(
(
Case
Case
(
Var
1
)
(
Var
0
)
(
InjL
Unit
)
(
InjL
Unit
)
(
(* try popping *)
(
(* try popping *)
If
If
(
CAS
(
CAS
(
st
.[
ren
(+
7
)])
(
st
.[
ren
(+
5
)])
(
Var
4
)
(
Var
2
)
(
Unfold
(
Snd
(
Var
0
)))
(
Unfold
(
Snd
(
Var
0
)))
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
5
)
(
Var
6
))
(* pop fails, we retry*)
)
)
)
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
(
App
(
Var
3
)
(
Var
4
))
(* pop fails, we retry*)
(
Load
(
Var
1
))
)
)
)
)
)
)
(
Load
st
.[
ren
(+
2
)])
).
).
Definition
FG_iter
(
f
:
expr
)
:
expr
:
=
Definition
FG_iter
(
f
:
expr
)
:
expr
:
=
Rec
Rec
(
Case
(
Load
(
Unfold
(
Var
1
)))
(
Case
(
Load
(
Unfold
(
Var
1
)))
Unit
Unit
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(* recursive_call *)
(
Seq
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
(
Var
1
)
(
Snd
(
Var
0
)))
(* recursive_call *)
)
)
).
).
Definition
FG_iterV
(
f
:
expr
)
:
val
:
=
Definition
FG_iterV
(
f
:
expr
)
:
val
:
=
RecV
RecV
(
Case
(
Load
(
Unfold
(
Var
1
)))
(
Case
(
Load
(
Unfold
(
Var
1
)))
Unit
Unit
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(* recursive_call *)
(
Seq
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
(
Var
1
)
(
Snd
(
Var
0
)))
(* recursive_call *)
)
)
).
).
Definition
FG_read_iter
(
st
:
expr
)
:
expr
:
=
Definition
FG_read_iter
(
st
:
expr
)
:
expr
:
=
Rec
(
App
(
FG_iter
(
Var
1
))
(
Fold
(
Load
st
.[
ren
(+
2
)]))).
Lam
(
App
(
FG_iter
(
Var
0
))
(
Fold
(
Load
st
.[
ren
(+
1
)]))).
Definition
FG_stack_body
(
st
:
expr
)
:
expr
:
=
Definition
FG_stack_body
(
st
:
expr
)
:
expr
:
=
Pair
(
Pair
(
FG_push
st
)
(
FG_pop
st
))
(
FG_read_iter
st
).
Pair
(
Pair
(
FG_push
st
)
(
FG_pop
st
))
(
FG_read_iter
st
).
...
@@ -116,34 +112,19 @@ Section FG_stack.
...
@@ -116,34 +112,19 @@ Section FG_stack.
(* Fine-grained push *)
(* Fine-grained push *)
Lemma
FG_push_folding
(
st
:
expr
)
:
Lemma
FG_push_folding
(
st
:
expr
)
:
FG_push
st
=
FG_push
st
=
Rec
(
App
(
Rec
Rec
(* try push *)
(
LetIn
(
If
(
CAS
(
st
.[
ren
(+
4
)])
(
Var
1
)
(
Load
st
.[
ren
(+
2
)]
)
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Fold
(
Var
1
))))
)
(* try push *
)
)
(
If
(
CAS
(
st
.[
ren
(+
3
)])
(
Var
0
)
Unit
(* push succeeds we return unit *
)
(
Alloc
(
InjR
(
Pair
(
Var
2
)
(
Fold
(
Var
0
))))
)
(
App
(
Var
2
)
(
Var
3
))
(* push fails, we try again *
)
)
)
Unit
(* push succeeds we return unit *
)
)
(
App
(
Var
1
)
(
Var
2
))
(* push fails, we try again *
)
(
Load
st
.[
ren
(+
2
)])
(* read the stack pointer *
)
)
).
).
Proof
.
trivial
.
Qed
.
Proof
.
trivial
.
Qed
.
Section
FG_push_type
.
Lemma
FG_push_type
st
Γ
τ
:
typed
Γ
st
(
Tref
(
FG_Stack_Ref_Type
τ
))
→
typed
Γ
(
FG_push
st
)
(
TArrow
τ
TUnit
).
Proof
.
intros
HTst
.
repeat
match
goal
with
|-
typed
_
_
_
=>
econstructor
;
eauto
end
;
repeat
econstructor
;
eauto
.
-
eapply
(
context_weakening
[
_;
_;
_;
_
])
;
eauto
.
-
by
asimpl
.
-
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
Qed
.
End
FG_push_type
.
Lemma
FG_push_to_val
st
:
to_val
(
FG_push
st
)
=
Some
(
FG_pushV
st
).
Lemma
FG_push_to_val
st
:
to_val
(
FG_push
st
)
=
Some
(
FG_pushV
st
).
Proof
.
trivial
.
Qed
.
Proof
.
trivial
.
Qed
.
...
@@ -151,83 +132,95 @@ Section FG_stack.
...
@@ -151,83 +132,95 @@ Section FG_stack.
Lemma
FG_push_of_val
st
:
of_val
(
FG_pushV
st
)
=
FG_push
st
.
Lemma
FG_push_of_val
st
:
of_val
(
FG_pushV
st
)
=
FG_push
st
.
Proof
.
trivial
.
Qed
.
Proof
.
trivial
.
Qed
.
Global
Typeclasses
Opaque
FG_pushV
.
Global
Opaque
FG_pushV
.
Global
Opaque
FG_pushV
.
Lemma
FG_push_closed
(
st
:
expr
)
:
Lemma
FG_push_type
st
Γ
τ
:
(
∀
f
,
st
.[
f
]
=
st
)
→
∀
f
,
(
FG_push
st
).[
f
]
=
FG_push
st
.
typed
Γ
st
(
Tref
(
FG_Stack_Ref_Type
τ
))
→
Proof
.
intros
H
f
.
asimpl
.
unfold
FG_push
.
rewrite
?H
;
trivial
.
Qed
.
typed
Γ
(
FG_push
st
)
(
TArrow
τ
TUnit
).
Proof
.
intros
?.
do
2
econstructor
.
{
econstructor
;
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
}
econstructor
;
eauto
using
typed
.
econstructor
;
eauto
using
typed
;
repeat
econstructor
.
-
eapply
(
context_weakening
[
_;
_;
_
])
;
eauto
.
-
by
asimpl
.
Qed
.
Lemma
FG_push_subst
(
st
:
expr
)
f
:
(
FG_push
st
).[
f
]
=
FG_push
st
.[
f
].
Lemma
FG_push_subst
(
st
:
expr
)
f
:
Proof
.
unfold
FG_push
.
by
asimpl
.
Qed
.
(
FG_push
st
).[
f
]
=
FG_push
st
.[
f
].
Proof
.
by
rewrite
/
FG_push
;
asimpl
.
Qed
.
Hint
Rewrite
FG_push_subst
:
autosubst
.
Global
Typeclasses
Opaque
FG_push
.
Global
Opaque
FG_push
.
Global
Opaque
FG_push
.
(* Fine-grained push *)
(* Fine-grained push *)
Lemma
FG_pop_folding
(
st
:
expr
)
:
Lemma
FG_pop_folding
(
st
:
expr
)
:
FG_pop
st
=
FG_pop
st
=
Rec
(
App
(
Rec
Rec
(
App
(
LetIn
(
Rec
(
Load
st
.[
ren
(+
2
)])
(
(
LetIn
Case
(
Var
1
)
(
Load
(
Var
0
))
(
InjL
Unit
)
(
Case
(
(* try popping *)
(
Var
0
)
If
(
InjL
Unit
)
(
CAS
(
(* try popping *)
(
st
.[
ren
(+
7
)])
If
(
Var
4
)
(
CAS
(
Unfold
(
Snd
(
Var
0
)))
(
st
.[
ren
(+
5
)])
)
(
Var
2
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
Unfold
(
Snd
(
Var
0
)))
(
App
(
Var
5
)
(
Var
6
))
(* pop fails, we retry*)
)
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
)
(
App
(
Var
3
)
(
Var
4
))
(* pop fails, we retry*)
)
)
(
)
(
Load
(
Var
1
))
)
)
).
)
)
(
Load
st
.[
ren
(+
2
)])
).
Proof
.
trivial
.
Qed
.
Proof
.
trivial
.
Qed
.
Section
FG_pop_type
.
Lemma
FG_pop_type
st
Γ
τ
:
typed
Γ
st
(
Tref
(
FG_Stack_Ref_Type
τ
))
→
typed
Γ
(
FG_pop
st
)
(
TArrow
TUnit
(
TSum
TUnit
τ
)).