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
Gaurav Parthasarathy
examples_rdcss_old
Commits
76c60e8f
Commit
76c60e8f
authored
Dec 10, 2018
by
Amin Timany
Browse files
Simplify stack programs
parent
f72453a9
Changes
4
Expand all
Hide whitespace changes
Inline
Side-by-side
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
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
τ
:
=
TRec
(
Tref
(
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
)))).
...
...
@@ -7,103 +7,99 @@ Definition FG_Stack_Ref_Type τ :=
Tref
(
TSum
TUnit
(
TProd
τ
(
FG_StackType
τ
))).
Definition
FG_push
(
st
:
expr
)
:
expr
:
=
Rec
(
App
(
Rec
(* try push *)
(
If
(
CAS
(
st
.[
ren
(+
4
)])
(
Var
1
)
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Fold
(
Var
1
)))))
)
Unit
(* push succeeds we return unit *)
(
App
(
Var
2
)
(
Var
3
))
(* push fails, we try again *)
)
Rec
(
LetIn
(
Load
st
.[
ren
(+
2
)])
(* try push *)
(
If
(
CAS
(
st
.[
ren
(+
3
)])
(
Var
0
)
(
Alloc
(
InjR
(
Pair
(
Var
2
)
(
Fold
(
Var
0
)))))
)
(
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
:
=
RecV
(
App
(
Rec
(* try push *)
(
If
(
CAS
(
st
.[
ren
(+
4
)])
(
Var
1
)
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Fold
(
Var
1
)))))
)
Unit
(* push succeeds we return unit *)
(
App
(
Var
2
)
(
Var
3
))
(* push fails, we try again *)
)
RecV
(
LetIn
(
Load
st
.[
ren
(+
2
)])
(* try push *)
(
If
(
CAS
(
st
.[
ren
(+
3
)])
(
Var
0
)
(
Alloc
(
InjR
(
Pair
(
Var
2
)
(
Fold
(
Var
0
)))))
)
(
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
:
=
Rec
(
App
(
Rec
(
App
(
Rec
(
Case
(
Var
1
)
(
InjL
Unit
)
(
(* try popping *)
If
(
CAS
(
st
.[
ren
(+
7
)])
(
Var
4
)
(
Unfold
(
Snd
(
Var
0
)))
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
5
)
(
Var
6
))
(* pop fails, we retry*)
)
)
)
(
(
Load
(
Var
1
))
Rec
(
LetIn
(
Load
st
.[
ren
(+
2
)])
(
LetIn
(
Load
(
Var
0
))
(
Case
(
Var
0
)
(
InjL
Unit
)
(
(* try popping *)
If
(
CAS
(
st
.[
ren
(+
5
)])
(
Var
2
)
(
Unfold
(
Snd
(
Var
0
)))
)
)
)
(
Load
st
.[
ren
(+
2
)])
).
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
3
)
(
Var
4
))
(* pop fails, we retry*)
)
)
)
).
Definition
FG_popV
(
st
:
expr
)
:
val
:
=
RecV
(
App
(
Rec
(
App
(
Rec
(
Case
(
Var
1
)
(
InjL
Unit
)
(
(* try popping *)
If
(
CAS
(
st
.[
ren
(+
7
)])
(
Var
4
)
(
Unfold
(
Snd
(
Var
0
)))
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
5
)
(
Var
6
))
(* pop fails, we retry*)
)
(
LetIn
(
Load
st
.[
ren
(+
2
)])
(
LetIn
(
Load
(
Var
0
))
(
Case
(
Var
0
)
(
InjL
Unit
)
(
(* try popping *)
If
(
CAS
(
st
.[
ren
(+
5
)])
(
Var
2
)
(
Unfold
(
Snd
(
Var
0
)))
)
)
(
(
Load
(
Var
1
))
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
3
)
(
Var
4
))
(* pop fails, we retry*)
)
)
)
(
Load
st
.[
ren
(+
2
)])
).
Definition
FG_iter
(
f
:
expr
)
:
expr
:
=
Rec
(
Case
(
Load
(
Unfold
(
Var
1
)))
Unit
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(* recursive_call *)
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
Seq
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
(
Var
1
)
(
Snd
(
Var
0
)))
(* recursive_call *)
)
).
Definition
FG_iterV
(
f
:
expr
)
:
val
:
=
RecV
(
Case
(
Load
(
Unfold
(
Var
1
)))
Unit
(
App
(
Rec
(
App
(
Var
3
)
(
Snd
(
Var
2
))))
(* recursive_call *)
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
Seq
(
App
f
.[
ren
(+
3
)]
(
Fst
(
Var
0
)))
(
App
(
Var
1
)
(
Snd
(
Var
0
)))
(* recursive_call *)
)
).
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
:
=
Pair
(
Pair
(
FG_push
st
)
(
FG_pop
st
))
(
FG_read_iter
st
).
...
...
@@ -116,34 +112,19 @@ Section FG_stack.
(* Fine-grained push *)
Lemma
FG_push_folding
(
st
:
expr
)
:
FG_push
st
=
Rec
(
App
(
Rec
(* try push *)
(
If
(
CAS
(
st
.[
ren
(+
4
)])
(
Var
1
)
(
Alloc
(
InjR
(
Pair
(
Var
3
)
(
Fold
(
Var
1
))))
)
)
Unit
(* push succeeds we return unit *
)
(
App
(
Var
2
)
(
Var
3
))
(* push fails, we try again *
)
)
)
(
Load
st
.[
ren
(+
2
)])
(* read the stack pointer *
)
).
Rec
(
LetIn
(
Load
st
.[
ren
(+
2
)])
(* try push *
)
(
If
(
CAS
(
st
.[
ren
(+
3
)])
(
Var
0
)
(
Alloc
(
InjR
(
Pair
(
Var
2
)
(
Fold
(
Var
0
))))
)
)
Unit
(* push succeeds we return unit *
)
(
App
(
Var
1
)
(
Var
2
))
(* push fails, we try again *
)
)
).
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
).
Proof
.
trivial
.
Qed
.
...
...
@@ -151,83 +132,95 @@ Section FG_stack.
Lemma
FG_push_of_val
st
:
of_val
(
FG_pushV
st
)
=
FG_push
st
.
Proof
.
trivial
.
Qed
.
Global
Typeclasses
Opaque
FG_pushV
.
Global
Opaque
FG_pushV
.
Lemma
FG_push_closed
(
st
:
expr
)
:
(
∀
f
,
st
.[
f
]
=
st
)
→
∀
f
,
(
FG_push
st
).[
f
]
=
FG_push
st
.
Proof
.
intros
H
f
.
asimpl
.
unfold
FG_push
.
rewrite
?H
;
trivial
.
Qed
.
Lemma
FG_push_type
st
Γ
τ
:
typed
Γ
st
(
Tref
(
FG_Stack_Ref_Type
τ
))
→
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
].
Proof
.
unfold
FG_push
.
by
asimpl
.
Qed
.
Lemma
FG_push_subst
(
st
:
expr
)
f
:
(
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
.
(* Fine-grained push *)
Lemma
FG_pop_folding
(
st
:
expr
)
:
FG_pop
st
=
Rec
(
App
(
Rec
(
App
(
Rec
(
Case
(
Var
1
)
(
InjL
Unit
)
(
(* try popping *)
If
(
CAS
(
st
.[
ren
(+
7
)])
(
Var
4
)
(
Unfold
(
Snd
(
Var
0
)))
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
5
)
(
Var
6
))
(* pop fails, we retry*)
)
)
)
(
(
Load
(
Var
1
))
)
)
)
(
Load
st
.[
ren
(+
2
)])
).
Rec
(
LetIn
(
Load
st
.[
ren
(+
2
)])
(
LetIn
(
Load
(
Var
0
))
(
Case
(
Var
0
)
(
InjL
Unit
)
(
(* try popping *)
If
(
CAS
(
st
.[
ren
(+
5
)])
(
Var
2
)
(
Unfold
(
Snd
(
Var
0
)))
)
(
InjR
(
Fst
(
Var
0
)))
(* pop succeeds *)
(
App
(
Var
3
)
(
Var
4
))
(* pop fails, we retry*)
)
)
)
).
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
τ
)).
Proof
.
replace
(
FG_Stack_Ref_Type
τ
)
with
(
Tref
(
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
)))).[
FG_StackType
τ
/]
;
last
first
.
{
by
asimpl
.
}
intros
HTst
.
repeat
match
goal
with
|-
typed
_
_
_
=>
econstructor
;
eauto
end
;
repeat
econstructor
;
eauto
;
last
first
.
-
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
-
by
asimpl
.
-
eapply
(
context_weakening
[
_;
_;
_;
_;
_;
_;
_
])
;
eauto
.
-
econstructor
.
Qed
.
End
FG_pop_type
.
Lemma
FG_pop_to_val
st
:
to_val
(
FG_pop
st
)
=
Some
(
FG_popV
st
).
Proof
.
trivial
.
Qed
.
Lemma
FG_pop_of_val
st
:
of_val
(
FG_popV
st
)
=
FG_pop
st
.
Proof
.
trivial
.
Qed
.
Global
Typeclasses
Opaque
FG_popV
.
Global
Opaque
FG_popV
.
Lemma
FG_pop_closed
(
st
:
expr
)
:
(
∀
f
,
st
.[
f
]
=
st
)
→
∀
f
,
(
FG_pop
st
).[
f
]
=
FG_pop
st
.
Proof
.
intros
H
f
.
asimpl
.
unfold
FG_pop
.
rewrite
?H
;
trivial
.
Qed
.
Lemma
FG_pop_type
st
Γ
τ
:
typed
Γ
st
(
Tref
(
FG_Stack_Ref_Type
τ
))
→
typed
Γ
(
FG_pop
st
)
(
TArrow
TUnit
(
TSum
TUnit
τ
)).
Proof
.
replace
(
FG_Stack_Ref_Type
τ
)
with
(
Tref
(
TSum
TUnit
(
TProd
τ
.[
ren
(+
1
)]
(
TVar
0
)))).[
FG_StackType
τ
/]
;
last
first
.
{
by
asimpl
.
}
intros
?.
do
2
econstructor
.
{
econstructor
.
eapply
(
context_weakening
[
_;
_
])
;
eauto
.
}
econstructor
.
{
repeat
econstructor
.
}
econstructor
.
{
repeat
econstructor
.
}
{
repeat
econstructor
.
}
econstructor
.
-
econstructor
;
[|
eapply
(
context_weakening
[
_;
_;
_;
_;
_
])
;
eauto
|
|]
;
repeat
econstructor
.
-
by
repeat
econstructor
;
asimpl
.
-
repeat
econstructor
.
Qed
.
Lemma
FG_pop_subst
(
st
:
expr
)
f
:
(
FG_pop
st
).[
f
]
=
FG_pop
st
.[
f
].
Proof
.
unfold
FG_pop
.
by
asimpl
.
Qed
.
Hint
Rewrite
FG_pop_subst
:
autosubst
.
Global
Typeclasses
Opaque
FG_pop
.
Global
Opaque
FG_pop
.
(* Fine-grained iter *)
...
...
@@ -236,8 +229,9 @@ Section FG_stack.
Rec