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
88c9a1e0
Commit
88c9a1e0
authored
Jul 02, 2016
by
Robbert Krebbers
Browse files
Make logical relations look like logical relations.
parent
b98e54c9
Changes
21
Expand all
Hide whitespace changes
Inline
Side-by-side
F_mu/fundamental.v
View file @
88c9a1e0
...
...
@@ -16,20 +16,14 @@ Section typed_interp.
Local
Ltac
value_case
:=
iApply
wp_value
;
cbn
;
rewrite
?
to_of_val
;
trivial
.
Lemma
typed_interp
Δ
Γ
vs
e
τ
(
H
Δ
:
ctx_PersistentP
Δ
)
:
Γ
⊢ₜ
e
:
τ
→
length
Γ
=
length
vs
→
[
∧
]
zip_with
(
λ
τ
,
interp
τ
Δ
)
Γ
vs
⊢
WP
e
.[
env_subst
vs
]
{{
interp
τ
Δ
}}
.
Theorem
fundamental
Δ
Γ
vs
e
τ
(
H
Δ
:
env_PersistentP
Δ
)
:
Γ
⊢ₜ
e
:
τ
→
⟦
Γ
⟧
*
Δ
vs
⊢
⟦
τ
⟧ₑ
Δ
e
.[
env_subst
vs
].
Proof
.
intros
Htyped
.
revert
Δ
H
Δ
vs
.
induction
Htyped
;
iIntros
{
Δ
H
Δ
vs
H
len
}
"#HΓ"
;
cbn
.
intros
Htyped
.
revert
Δ
vs
H
Δ
.
induction
Htyped
;
iIntros
{
Δ
vs
H
Δ
}
"#HΓ"
;
cbn
.
-
(
*
var
*
)
destruct
(
lookup_lt_is_Some_2
vs
x
)
as
[
v
Hv
].
{
by
rewrite
-
Hlen
;
apply
lookup_lt_Some
with
τ
.
}
rewrite
/
env_subst
Hv
;
value_case
.
iApply
(
big_and_elem_of
with
"HΓ"
).
apply
elem_of_list_lookup_2
with
x
.
rewrite
lookup_zip_with
;
by
simplify_option_eq
.
iDestruct
(
interp_env_Some_l
with
"HΓ"
)
as
{
v
}
"[% ?]"
;
first
done
.
rewrite
/
env_subst
.
simplify_option_eq
.
by
value_case
.
-
(
*
unit
*
)
value_case
.
-
(
*
pair
*
)
smart_wp_bind
(
PairLCtx
e2
.[
env_subst
vs
])
v
"# Hv"
IHHtyped1
.
...
...
@@ -51,47 +45,43 @@ Section typed_interp.
value_case
;
eauto
.
-
(
*
case
*
)
smart_wp_bind
(
CaseCtx
_
_
)
v
"#Hv"
IHHtyped1
;
cbn
.
iDestruct
"Hv"
as
"[Hv|Hv]"
;
iDestruct
"Hv"
as
{
w
}
"[% Hw]"
;
subst
.
+
iApply
wp_case_inl
;
auto
1
using
to_of_val
;
asimpl
.
specialize
(
IHHtyped2
Δ
H
Δ
(
w
::
vs
)).
erewrite
<-
?
typed_subst_head_simpl
in
*
by
(
cbn
;
eauto
).
iNext
;
iApply
IHHtyped2
;
cbn
;
auto
.
+
iApply
wp_case_inr
;
auto
1
using
to_of_val
;
asimpl
.
specialize
(
IHHtyped3
Δ
H
Δ
(
w
::
vs
)).
erewrite
<-
?
typed_subst_head_simpl
in
*
by
(
cbn
;
eauto
).
iNext
;
iApply
IHHtyped3
;
cbn
;
auto
.
iDestruct
(
interp_env_length
with
"HΓ"
)
as
%?
.
iDestruct
"Hv"
as
"[Hv|Hv]"
;
iDestruct
"Hv"
as
{
w
}
"[% Hw]"
;
simplify_eq
/=
.
+
iApply
wp_case_inl
;
auto
1
using
to_of_val
;
asimpl
.
iNext
.
erewrite
typed_subst_head_simpl
by
naive_solver
.
iApply
(
IHHtyped2
Δ
(
w
::
vs
)).
iApply
interp_env_cons
;
auto
.
+
iApply
wp_case_inr
;
auto
1
using
to_of_val
;
asimpl
.
iNext
.
erewrite
typed_subst_head_simpl
by
naive_solver
.
iApply
(
IHHtyped3
Δ
(
w
::
vs
)).
iApply
interp_env_cons
;
auto
.
-
(
*
lam
*
)
value_case
;
iAlways
;
iIntros
{
w
}
"#Hw"
.
iApply
wp_lam
;
auto
1
using
to_of_val
.
asimpl
;
erewrite
typed_subst_head_simpl
;
[
|
eauto
|
cbn
];
eauto
.
iNext
;
iApply
(
IHHtyped
Δ
H
Δ
(
w
::
vs
));
cbn
;
auto
.
iDestruct
(
interp_env_length
with
"HΓ"
)
as
%?
.
iApply
wp_lam
;
auto
1
using
to_of_val
.
iNext
.
asimpl
.
erewrite
typed_subst_head_simpl
by
naive_solver
.
iApply
(
IHHtyped
Δ
(
w
::
vs
)).
iApply
interp_env_cons
;
auto
.
-
(
*
app
*
)
smart_wp_bind
(
AppLCtx
(
e2
.[
env_subst
vs
]))
v
"#Hv"
IHHtyped1
.
smart_wp_bind
(
AppRCtx
v
)
w
"#Hw"
IHHtyped2
.
iApply
wp_mono
;
[
|
iApply
"Hv"
];
auto
.
-
(
*
TLam
*
)
value_case
.
iAlways
;
iIntros
{
τ
i
}
"%"
.
iApply
wp_TLam
;
iNext
.
simpl
in
*
.
iApply
(
IHHtyped
(
τ
i
::
Δ
)).
by
rewrite
fmap_length
.
rewrite
zip_with_fmap_l
.
by
iApply
context_interp_ren_S
.
iAlways
;
iIntros
{
τ
i
}
"%"
.
iApply
wp_TLam
;
iNext
.
iApply
IHHtyped
.
by
iApply
interp_env_ren
.
-
(
*
TApp
*
)
smart_wp_bind
TAppCtx
v
"#Hv"
IHHtyped
;
cbn
.
iApply
wp_wand_r
;
iSplitL
;
[
iApply
(
"Hv"
$
!
(
interp
τ'
Δ
));
iPureIntro
;
apply
_
|
].
iApply
wp_wand_r
;
iSplitL
;
[
iApply
(
"Hv"
$
!
(
⟦
τ'
⟧
Δ
));
iPureIntro
;
apply
_
|
].
iIntros
{
w
}
"?"
.
by
rewrite
interp_subst
.
-
(
*
Fold
*
)
rewrite
map_length
in
IHHtyped
.
iApply
(
@
wp_bind
_
_
_
[
FoldCtx
]).
iApply
wp_wand_l
.
iSplitL
;
[
|
iApply
(
IHHtyped
(
interp
(
TRec
τ
)
Δ
::
Δ
));
trivial
].
+
iIntros
{
v
}
"#Hv"
.
value_case
.
change
(
fixpoint
_
)
with
(
interp
(
TRec
τ
)
Δ
)
at
1
;
trivial
.
rewrite
fixpoint_unfold
/=
.
iAlways
;
eauto
10.
+
rewrite
zip_with_fmap_l
.
by
iApply
context_interp_ren_S
.
iApply
(
@
wp_bind
_
_
_
[
FoldCtx
]);
iApply
wp_wand_l
;
iSplitL
;
[
|
iApply
(
IHHtyped
Δ
vs
);
auto
].
iIntros
{
v
}
"#Hv"
.
value_case
.
rewrite
/=
-
interp_subst
fixpoint_unfold
/=
.
iAlways
;
eauto
.
-
(
*
Unfold
*
)
iApply
(
@
wp_bind
_
_
_
[
UnfoldCtx
]);
iApply
wp_wand_l
;
iSplitL
;
[
|
iApply
IHHtyped
;
trivial
].
iIntros
{
v
}
"#Hv"
.
rewrite
/=
fixpoint_unfold
.
change
(
fixpoint
_
)
with
(
interp
(
TRec
τ
)
Δ
);
simpl
.
change
(
fixpoint
_
)
with
(
⟦
TRec
τ
⟧
Δ
);
simpl
.
iDestruct
"Hv"
as
{
w
}
"#[% Hw]"
;
subst
.
iApply
wp_Fold
;
cbn
;
auto
using
to_of_val
.
by
rewrite
-
interp_subst
.
...
...
F_mu/logrel.v
View file @
88c9a1e0
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris_logrel
.
F_mu
Require
Export
lang
typing
.
Import
uPred
.
...
...
@@ -65,29 +66,39 @@ Section logrel.
|
TForall
τ'
=>
interp_forall
(
interp
τ'
)
|
TRec
τ'
=>
interp_rec
(
interp
τ'
)
end
%
I
.
Notation
"⟦ τ ⟧"
:=
(
interp
τ
).
Class
ctx_PersistentP
Δ
:=
Definition
interp_env
(
Γ
:
list
type
)
(
Δ
:
listC
D
)
(
vs
:
list
val
)
:
iProp
lang
Σ
:=
(
length
Γ
=
length
vs
∧
[
∧
]
zip_with
(
λ
τ
,
⟦
τ
⟧
Δ
)
Γ
vs
)
%
I
.
Notation
"⟦ Γ ⟧*"
:=
(
interp_env
Γ
).
Definition
interp_expr
(
τ
:
type
)
(
Δ
:
listC
D
)
(
e
:
expr
)
:
iProp
lang
Σ
:=
WP
e
{{
⟦
τ
⟧
Δ
}}%
I
.
Class
env_PersistentP
Δ
:=
ctx_persistentP
:
Forall
(
λ
τ
i
,
∀
v
,
PersistentP
(
τ
i
v
))
Δ
.
Global
Instance
ctx_persistent_nil
:
ctx
_PersistentP
[].
Global
Instance
ctx_persistent_nil
:
env
_PersistentP
[].
Proof
.
by
constructor
.
Qed
.
Global
Instance
ctx_persistent_cons
τ
i
Δ
:
(
∀
v
,
PersistentP
(
τ
i
v
))
→
ctx
_PersistentP
Δ
→
ctx
_PersistentP
(
τ
i
::
Δ
).
(
∀
v
,
PersistentP
(
τ
i
v
))
→
env
_PersistentP
Δ
→
env
_PersistentP
(
τ
i
::
Δ
).
Proof
.
by
constructor
.
Qed
.
Global
Instance
ctx_persistent_lookup
Δ
x
v
:
ctx
_PersistentP
Δ
→
PersistentP
(
ctx_lookup
x
Δ
v
).
env
_PersistentP
Δ
→
PersistentP
(
ctx_lookup
x
Δ
v
).
Proof
.
intros
H
Δ
;
revert
x
;
induction
H
Δ
=>-
[
|?
]
/=
;
apply
_.
Qed
.
Global
Instance
interp_var_persistent
τ
Δ
v
:
ctx_PersistentP
Δ
→
PersistentP
(
interp
τ
Δ
v
).
Global
Instance
interp_persistent
τ
Δ
v
:
env_PersistentP
Δ
→
PersistentP
(
⟦
τ
⟧
Δ
v
).
Proof
.
revert
v
Δ
;
induction
τ
=>
v
Δ
H
Δ
;
simpl
;
try
apply
_.
rewrite
/
PersistentP
/
interp_rec
fixpoint_unfold
/
interp_rec1
/=
.
by
apply
always_intro
'
.
Qed
.
Global
Instance
interp_env_persistent
Γ
Δ
vs
:
env_PersistentP
Δ
→
PersistentP
(
⟦
Γ
⟧
*
Δ
vs
)
:=
_.
Lemma
interp_weaken
Δ
1
Π
Δ
2
τ
:
interp
τ
.[
iter
(
length
Δ
1
)
up
(
ren
(
+
length
Π
))]
(
Δ
1
++
Π
++
Δ
2
)
≡
interp
τ
(
Δ
1
++
Δ
2
).
⟦
τ
.[
iter
(
length
Δ
1
)
up
(
ren
(
+
length
Π
))]
⟧
(
Δ
1
++
Π
++
Δ
2
)
≡
⟦
τ
⟧
(
Δ
1
++
Δ
2
).
Proof
.
revert
Δ
1
Π
Δ
2.
induction
τ
=>
Δ
1
Π
Δ
2
;
simpl
;
auto
.
-
intros
w
;
simpl
;
properness
;
auto
.
by
apply
IH
τ
1.
by
apply
IH
τ
2.
...
...
@@ -103,8 +114,8 @@ Section logrel.
Qed
.
Lemma
interp_subst_up
Δ
1
Δ
2
τ
τ'
:
interp
τ
(
Δ
1
++
interp
τ'
Δ
2
::
Δ
2
)
≡
interp
τ
.[
iter
(
length
Δ
1
)
up
(
τ'
.
:
ids
)]
(
Δ
1
++
Δ
2
).
⟦
τ
⟧
(
Δ
1
++
interp
τ'
Δ
2
::
Δ
2
)
≡
⟦
τ
.[
iter
(
length
Δ
1
)
up
(
τ'
.
:
ids
)]
⟧
(
Δ
1
++
Δ
2
).
Proof
.
revert
Δ
1
Δ
2
;
induction
τ
=>
Δ
1
Δ
2
;
simpl
.
-
done
.
...
...
@@ -124,15 +135,41 @@ Section logrel.
-
intros
w
;
simpl
;
properness
;
auto
.
apply
(
IH
τ
(
_
::
_
)).
Qed
.
Lemma
interp_subst
Δ
2
τ
τ'
:
interp
τ
(
interp
τ'
Δ
2
::
Δ
2
)
≡
interp
τ
.[
τ'
/
]
Δ
2.
Lemma
interp_subst
Δ
2
τ
τ'
:
⟦
τ
⟧
(
⟦
τ'
⟧
Δ
2
::
Δ
2
)
≡
⟦
τ
.[
τ'
/
]
⟧
Δ
2.
Proof
.
apply
(
interp_subst_up
[]).
Qed
.
Lemma
context_interp_ren_S
Δ
(
Γ
:
list
type
)
(
vs
:
list
val
)
τ
i
:
([
∧
]
zip_with
(
λ
τ
,
interp
τ
Δ
)
Γ
vs
)
⊣⊢
([
∧
]
zip_with
(
λ
τ
,
interp
τ
.[
ren
(
+
1
)]
(
τ
i
::
Δ
))
Γ
vs
).
Lemma
interp_env_length
Δ
Γ
vs
:
⟦
Γ
⟧
*
Δ
vs
⊢
length
Γ
=
length
vs
.
Proof
.
by
iIntros
"[% ?]"
.
Qed
.
Lemma
interp_env_Some_l
Δ
Γ
vs
x
τ
:
Γ
!!
x
=
Some
τ
→
⟦
Γ
⟧
*
Δ
vs
⊢
∃
v
,
vs
!!
x
=
Some
v
∧
⟦
τ
⟧
Δ
v
.
Proof
.
revert
Δ
vs
τ
i
;
induction
Γ
=>
Δ
[
|
v
vs
]
τ
i
;
simpl
;
auto
.
apply
and_proper
;
auto
.
symmetry
.
apply
(
interp_weaken
[]
[
τ
i
]
Δ
).
iIntros
{?}
"[Hlen HΓ]"
;
iDestruct
"Hlen"
as
%
Hlen
.
destruct
(
lookup_lt_is_Some_2
vs
x
)
as
[
v
Hv
].
{
by
rewrite
-
Hlen
;
apply
lookup_lt_Some
with
τ
.
}
iExists
v
;
iSplit
.
done
.
iApply
(
big_and_elem_of
with
"HΓ"
).
apply
elem_of_list_lookup_2
with
x
.
rewrite
lookup_zip_with
;
by
simplify_option_eq
.
Qed
.
Lemma
interp_env_nil
Δ
:
True
⊢
⟦
[]
⟧
*
Δ
[].
Proof
.
iIntros
""
;
iSplit
;
auto
.
Qed
.
Lemma
interp_env_cons
Δ
Γ
vs
τ
v
:
⟦
τ
::
Γ
⟧
*
Δ
(
v
::
vs
)
⊣⊢
⟦
τ
⟧
Δ
v
∧
⟦
Γ
⟧
*
Δ
vs
.
Proof
.
rewrite
/
interp_env
/=
(
assoc
_
(
⟦
_
⟧
_
_
))
-
(
comm
_
(
_
=
_
)
%
I
)
-
assoc
.
by
apply
and_proper
;
[
apply
pure_proper
;
omega
|
].
Qed
.
Lemma
interp_env_ren
Δ
(
Γ
:
list
type
)
(
vs
:
list
val
)
τ
i
:
⟦
subst
(
ren
(
+
1
))
<
$
>
Γ
⟧
*
(
τ
i
::
Δ
)
vs
⊣⊢
⟦
Γ
⟧
*
Δ
vs
.
Proof
.
apply
and_proper
;
[
apply
pure_proper
;
by
rewrite
fmap_length
|
].
revert
Δ
vs
τ
i
;
induction
Γ
=>
Δ
[
|
v
vs
]
τ
i
;
csimpl
;
auto
.
apply
and_proper
;
auto
.
apply
(
interp_weaken
[]
[
τ
i
]
Δ
).
Qed
.
End
logrel
.
Notation
"⟦ τ ⟧"
:=
(
interp
τ
).
Notation
"⟦ τ ⟧ₑ"
:=
(
interp_expr
τ
).
Notation
"⟦ Γ ⟧*"
:=
(
interp_env
Γ
).
F_mu/soundness.v
View file @
88c9a1e0
...
...
@@ -9,7 +9,7 @@ Section soundness.
[]
⊢ₜ
e
:
τ
→
True
⊢
WP
e
{{
@
interp
(
globalF
Σ
)
τ
[]
}}
.
Proof
.
iIntros
{
H
}
""
.
rewrite
-
(
empty_env_subst
e
).
by
iApply
(
@
typed_interp
_
_
_
[])
.
iApply
(
@
fundamental
_
_
_
[]);
eauto
.
by
iApply
interp_env_nil
.
Qed
.
Theorem
soundness
e
τ
e
'
thp
:
...
...
F_mu/typing.v
View file @
88c9a1e0
...
...
@@ -32,7 +32,7 @@ Inductive typed (Γ : list type) : expr → type → Prop :=
|
App_typed
e1
e2
τ
1
τ
2
:
Γ
⊢ₜ
e1
:
TArrow
τ
1
τ
2
→
Γ
⊢ₜ
e2
:
τ
1
→
Γ
⊢ₜ
App
e1
e2
:
τ
2
|
TLam_typed
e
τ
:
subst
(
ren
(
+
1
))
<
$
>
Γ
⊢ₜ
e
:
τ
→
Γ
⊢ₜ
TLam
e
:
TForall
τ
|
TApp_typed
e
τ
τ'
:
Γ
⊢ₜ
e
:
TForall
τ
→
Γ
⊢ₜ
TApp
e
:
τ
.[
τ'
/
]
|
TFold
e
τ
:
subst
(
ren
(
+
1
))
<
$
>
Γ
⊢ₜ
e
:
τ
→
Γ
⊢ₜ
Fold
e
:
TRec
τ
|
TFold
e
τ
:
Γ
⊢ₜ
e
:
τ
.[
TRec
τ
/
]
→
Γ
⊢ₜ
Fold
e
:
TRec
τ
|
TUnfold
e
τ
:
Γ
⊢ₜ
e
:
TRec
τ
→
Γ
⊢ₜ
Unfold
e
:
τ
.[
TRec
τ
/
]
where
"Γ ⊢ₜ e : τ"
:=
(
typed
Γ
e
τ
).
...
...
F_mu_ref/fundamental.v
View file @
88c9a1e0
...
...
@@ -15,21 +15,15 @@ Section typed_interp.
Local
Ltac
value_case
:=
iApply
wp_value
;
[
cbn
;
rewrite
?
to_of_val
;
trivial
|
].
Lemma
typed_interp
Δ
Γ
vs
e
τ
(
H
Δ
:
ctx
_PersistentP
Δ
)
:
Theorem
fundamental
Δ
Γ
vs
e
τ
(
H
Δ
:
env
_PersistentP
Δ
)
:
Γ
⊢ₜ
e
:
τ
→
length
Γ
=
length
vs
→
heap_ctx
∧
[
∧
]
zip_with
(
λ
τ
,
interp
τ
Δ
)
Γ
vs
⊢
WP
e
.[
env_subst
vs
]
{{
interp
τ
Δ
}}
.
heap_ctx
∧
⟦
Γ
⟧
*
Δ
vs
⊢
⟦
τ
⟧ₑ
Δ
e
.[
env_subst
vs
].
Proof
.
intros
Htyped
;
revert
Δ
H
Δ
vs
.
induction
Htyped
;
iIntros
{
Δ
H
Δ
vs
H
len
}
"#[Hheap HΓ] /="
.
intros
Htyped
;
revert
Δ
vs
H
Δ
.
induction
Htyped
;
iIntros
{
Δ
vs
H
Δ
}
"#[Hheap HΓ] /="
.
-
(
*
var
*
)
destruct
(
lookup_lt_is_Some_2
vs
x
)
as
[
v
Hv
].
{
by
rewrite
-
Hlen
;
apply
lookup_lt_Some
with
τ
.
}
rewrite
/
env_subst
Hv
;
value_case
.
iApply
(
big_and_elem_of
with
"HΓ"
).
apply
elem_of_list_lookup_2
with
x
.
by
rewrite
lookup_zip_with
;
simplify_option_eq
.
iDestruct
(
interp_env_Some_l
with
"HΓ"
)
as
{
v
}
"[% ?]"
;
first
done
.
rewrite
/
env_subst
.
simplify_option_eq
.
by
value_case
.
-
(
*
unit
*
)
by
value_case
.
-
(
*
pair
*
)
smart_wp_bind
(
PairLCtx
e2
.[
env_subst
vs
])
v
"#Hv"
IHHtyped1
.
...
...
@@ -51,41 +45,38 @@ Section typed_interp.
value_case
;
eauto
.
-
(
*
case
*
)
smart_wp_bind
(
CaseCtx
_
_
)
v
"#Hv"
IHHtyped1
;
cbn
.
iDestruct
"Hv"
as
"[Hv|Hv]"
;
iDestruct
"Hv"
as
{
w
}
"[% Hw]"
;
subst
.
+
iApply
wp_case_inl
;
auto
1
using
to_of_val
;
asimpl
.
specialize
(
IHHtyped2
Δ
H
Δ
(
w
::
vs
)).
erewrite
<-
?
typed_subst_head_simpl
in
*
by
(
cbn
;
eauto
).
iNext
;
iApply
IHHtyped2
;
cbn
;
auto
.
+
iApply
wp_case_inr
;
auto
1
using
to_of_val
;
asimpl
.
specialize
(
IHHtyped3
Δ
H
Δ
(
w
::
vs
)).
erewrite
<-
?
typed_subst_head_simpl
in
*
by
(
cbn
;
eauto
).
iNext
;
iApply
IHHtyped3
;
cbn
;
auto
.
iDestruct
(
interp_env_length
with
"HΓ"
)
as
%?
.
iDestruct
"Hv"
as
"[Hv|Hv]"
;
iDestruct
"Hv"
as
{
w
}
"[% Hw]"
;
simplify_eq
/=
.
+
iApply
wp_case_inl
;
auto
1
using
to_of_val
;
asimpl
.
iNext
.
erewrite
typed_subst_head_simpl
by
naive_solver
.
iApply
(
IHHtyped2
Δ
(
w
::
vs
)).
iSplit
;
[
|
iApply
interp_env_cons
];
auto
.
+
iApply
wp_case_inr
;
auto
1
using
to_of_val
;
asimpl
.
iNext
.
erewrite
typed_subst_head_simpl
by
naive_solver
.
iApply
(
IHHtyped3
Δ
(
w
::
vs
)).
iSplit
;
[
|
iApply
interp_env_cons
];
auto
.
-
(
*
lam
*
)
value_case
;
iAlways
;
iIntros
{
w
}
"#Hw"
.
iApply
wp_lam
;
auto
1
using
to_of_val
.
asimpl
;
erewrite
typed_subst_head_simpl
;
[
|
eauto
|
cbn
];
eauto
.
iNext
;
iApply
(
IHHtyped
Δ
H
Δ
(
w
::
vs
));
cbn
;
auto
.
iDestruct
(
interp_env_length
with
"HΓ"
)
as
%?
.
iApply
wp_lam
;
auto
1
using
to_of_val
.
iNext
.
asimpl
.
erewrite
typed_subst_head_simpl
by
naive_solver
.
iApply
(
IHHtyped
Δ
(
w
::
vs
)).
iFrame
"Hheap"
.
iApply
interp_env_cons
;
auto
.
-
(
*
app
*
)
smart_wp_bind
(
AppLCtx
(
e2
.[
env_subst
vs
]))
v
"#Hv"
IHHtyped1
.
smart_wp_bind
(
AppRCtx
v
)
w
"#Hw"
IHHtyped2
.
iApply
wp_mono
;
[
|
iApply
"Hv"
];
auto
.
-
(
*
TLam
*
)
value_case
.
iAlways
;
iIntros
{
τ
i
}
"%"
.
iApply
wp_TLam
;
iNext
.
simpl
in
*
.
iApply
(
IHHtyped
(
τ
i
::
Δ
)).
by
rewrite
fmap_length
.
iFrame
"Hheap"
.
rewrite
zip_with_fmap_l
.
by
iApply
context_interp_ren_S
.
iAlways
;
iIntros
{
τ
i
}
"%"
.
iApply
wp_TLam
;
iNext
.
iApply
IHHtyped
.
iFrame
"Hheap"
.
by
iApply
interp_env_ren
.
-
(
*
TApp
*
)
smart_wp_bind
TAppCtx
v
"#Hv"
IHHtyped
;
cbn
.
iApply
wp_wand_r
;
iSplitL
;
[
iApply
(
"Hv"
$
!
(
interp
τ'
Δ
));
iPureIntro
;
apply
_
|
].
iIntros
{
w
}
"?"
.
by
rewrite
interp_subst
.
-
(
*
Fold
*
)
rewrite
map_length
in
IHHtyped
.
iApply
(
@
wp_bind
_
_
_
[
FoldCtx
]).
iApply
wp_wand_l
.
iSplitL
;
[
|
iApply
(
IHHtyped
(
interp
(
TRec
τ
)
Δ
::
Δ
));
trivial
].
+
iIntros
{
v
}
"#Hv"
.
value_case
.
rewrite
fixpoint_unfold
/=
.
iAlways
;
eauto
10.
+
iFrame
"Hheap"
.
rewrite
zip_with_fmap_l
.
by
iApply
context_interp_ren_S
.
iApply
(
@
wp_bind
_
_
_
[
FoldCtx
]);
iApply
wp_wand_l
;
iSplitL
;
[
|
iApply
(
IHHtyped
Δ
vs
);
auto
].
iIntros
{
v
}
"#Hv"
.
value_case
.
rewrite
/=
-
interp_subst
fixpoint_unfold
/=
.
iAlways
;
eauto
.
-
(
*
Unfold
*
)
iApply
(
@
wp_bind
_
_
_
[
UnfoldCtx
]);
iApply
wp_wand_l
;
iSplitL
;
[
|
iApply
IHHtyped
;
auto
].
...
...
F_mu_ref/logrel.v
View file @
88c9a1e0
From
iris
.
pr
elu
de
Require
Import
string
s
.
From
iris
.
pr
oofmo
de
Require
Import
tactic
s
.
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris_logrel
.
F_mu_ref
Require
Export
rules
typing
.
Import
uPred
.
...
...
@@ -78,29 +78,39 @@ Section logrel.
|
TRec
τ'
=>
interp_rec
(
interp
τ'
)
|
Tref
τ'
=>
interp_ref
(
interp
τ'
)
end
.
Notation
"⟦ τ ⟧"
:=
(
interp
τ
).
Class
ctx_PersistentP
Δ
:=
Definition
interp_env
(
Γ
:
list
type
)
(
Δ
:
listC
D
)
(
vs
:
list
val
)
:
iPropG
lang
Σ
:=
(
length
Γ
=
length
vs
∧
[
∧
]
zip_with
(
λ
τ
,
⟦
τ
⟧
Δ
)
Γ
vs
)
%
I
.
Notation
"⟦ Γ ⟧*"
:=
(
interp_env
Γ
).
Definition
interp_expr
(
τ
:
type
)
(
Δ
:
listC
D
)
(
e
:
expr
)
:
iPropG
lang
Σ
:=
WP
e
{{
⟦
τ
⟧
Δ
}}%
I
.
Class
env_PersistentP
Δ
:=
ctx_persistentP
:
Forall
(
λ
τ
i
,
∀
v
,
PersistentP
(
τ
i
v
))
Δ
.
Global
Instance
ctx_persistent_nil
:
ctx
_PersistentP
[].
Global
Instance
ctx_persistent_nil
:
env
_PersistentP
[].
Proof
.
by
constructor
.
Qed
.
Global
Instance
ctx_persistent_cons
τ
i
Δ
:
(
∀
v
,
PersistentP
(
τ
i
v
))
→
ctx
_PersistentP
Δ
→
ctx
_PersistentP
(
τ
i
::
Δ
).
(
∀
v
,
PersistentP
(
τ
i
v
))
→
env
_PersistentP
Δ
→
env
_PersistentP
(
τ
i
::
Δ
).
Proof
.
by
constructor
.
Qed
.
Global
Instance
ctx_persistent_lookup
Δ
x
v
:
ctx
_PersistentP
Δ
→
PersistentP
(
ctx_lookup
x
Δ
v
).
env
_PersistentP
Δ
→
PersistentP
(
ctx_lookup
x
Δ
v
).
Proof
.
intros
H
Δ
;
revert
x
;
induction
H
Δ
=>-
[
|?
]
/=
;
apply
_.
Qed
.
Global
Instance
interp_var_persistent
τ
Δ
v
:
ctx_PersistentP
Δ
→
PersistentP
(
interp
τ
Δ
v
).
Global
Instance
interp_persistent
τ
Δ
v
:
env_PersistentP
Δ
→
PersistentP
(
⟦
τ
⟧
Δ
v
).
Proof
.
revert
v
Δ
;
induction
τ
=>
v
Δ
H
Δ
;
simpl
;
try
apply
_.
rewrite
/
PersistentP
/
interp_rec
fixpoint_unfold
/
interp_rec1
/=
.
by
apply
always_intro
'
.
Qed
.
Global
Instance
interp_env_persistent
Γ
Δ
vs
:
env_PersistentP
Δ
→
PersistentP
(
⟦
Γ
⟧
*
Δ
vs
)
:=
_.
Lemma
interp_weaken
Δ
1
Π
Δ
2
τ
:
interp
τ
.[
iter
(
length
Δ
1
)
up
(
ren
(
+
length
Π
))]
(
Δ
1
++
Π
++
Δ
2
)
≡
interp
τ
(
Δ
1
++
Δ
2
).
⟦
τ
.[
iter
(
length
Δ
1
)
up
(
ren
(
+
length
Π
))]
⟧
(
Δ
1
++
Π
++
Δ
2
)
≡
⟦
τ
⟧
(
Δ
1
++
Δ
2
).
Proof
.
revert
Δ
1
Π
Δ
2.
induction
τ
=>
Δ
1
Π
Δ
2
;
simpl
;
auto
.
-
intros
w
;
simpl
;
properness
;
auto
.
by
apply
IH
τ
1.
by
apply
IH
τ
2.
...
...
@@ -117,8 +127,8 @@ Section logrel.
Qed
.
Lemma
interp_subst_up
Δ
1
Δ
2
τ
τ'
:
interp
τ
(
Δ
1
++
interp
τ'
Δ
2
::
Δ
2
)
≡
interp
τ
.[
iter
(
length
Δ
1
)
up
(
τ'
.
:
ids
)]
(
Δ
1
++
Δ
2
).
⟦
τ
⟧
(
Δ
1
++
interp
τ'
Δ
2
::
Δ
2
)
≡
⟦
τ
.[
iter
(
length
Δ
1
)
up
(
τ'
.
:
ids
)]
⟧
(
Δ
1
++
Δ
2
).
Proof
.
revert
Δ
1
Δ
2
;
induction
τ
=>
Δ
1
Δ
2
;
simpl
.
-
done
.
...
...
@@ -139,15 +149,42 @@ Section logrel.
-
intros
w
;
simpl
;
properness
;
auto
.
by
apply
IH
τ
.
Qed
.
Lemma
interp_subst
Δ
2
τ
τ'
:
interp
τ
(
interp
τ'
Δ
2
::
Δ
2
)
≡
interp
τ
.[
τ'
/
]
Δ
2.
Lemma
interp_subst
Δ
2
τ
τ'
:
⟦
τ
⟧
(
⟦
τ'
⟧
Δ
2
::
Δ
2
)
≡
⟦
τ
.[
τ'
/
]
⟧
Δ
2.
Proof
.
apply
(
interp_subst_up
[]).
Qed
.
Lemma
context_interp_ren_S
Δ
(
Γ
:
list
type
)
(
vs
:
list
val
)
τ
i
:
([
∧
]
zip_with
(
λ
τ
,
interp
τ
Δ
)
Γ
vs
)
⊣⊢
([
∧
]
zip_with
(
λ
τ
,
interp
τ
.[
ren
(
+
1
)]
(
τ
i
::
Δ
))
Γ
vs
).
Lemma
interp_env_length
Δ
Γ
vs
:
⟦
Γ
⟧
*
Δ
vs
⊢
length
Γ
=
length
vs
.
Proof
.
by
iIntros
"[% ?]"
.
Qed
.
Lemma
interp_env_Some_l
Δ
Γ
vs
x
τ
:
Γ
!!
x
=
Some
τ
→
⟦
Γ
⟧
*
Δ
vs
⊢
∃
v
,
vs
!!
x
=
Some
v
∧
⟦
τ
⟧
Δ
v
.
Proof
.
revert
Δ
vs
τ
i
;
induction
Γ
=>
Δ
[
|
v
vs
]
τ
i
;
simpl
;
auto
.
apply
and_proper
;
auto
.
symmetry
.
apply
(
interp_weaken
[]
[
τ
i
]
Δ
).
iIntros
{?}
"[Hlen HΓ]"
;
iDestruct
"Hlen"
as
%
Hlen
.
destruct
(
lookup_lt_is_Some_2
vs
x
)
as
[
v
Hv
].
{
by
rewrite
-
Hlen
;
apply
lookup_lt_Some
with
τ
.
}
iExists
v
;
iSplit
.
done
.
iApply
(
big_and_elem_of
with
"HΓ"
).
apply
elem_of_list_lookup_2
with
x
.
rewrite
lookup_zip_with
;
by
simplify_option_eq
.
Qed
.
Lemma
interp_env_nil
Δ
:
True
⊢
⟦
[]
⟧
*
Δ
[].
Proof
.
iIntros
""
;
iSplit
;
auto
.
Qed
.
Lemma
interp_env_cons
Δ
Γ
vs
τ
v
:
⟦
τ
::
Γ
⟧
*
Δ
(
v
::
vs
)
⊣⊢
⟦
τ
⟧
Δ
v
∧
⟦
Γ
⟧
*
Δ
vs
.
Proof
.
rewrite
/
interp_env
/=
(
assoc
_
(
⟦
_
⟧
_
_
))
-
(
comm
_
(
_
=
_
)
%
I
)
-
assoc
.
by
apply
and_proper
;
[
apply
pure_proper
;
omega
|
].
Qed
.
Lemma
interp_env_ren
Δ
(
Γ
:
list
type
)
(
vs
:
list
val
)
τ
i
:
⟦
subst
(
ren
(
+
1
))
<
$
>
Γ
⟧
*
(
τ
i
::
Δ
)
vs
⊣⊢
⟦
Γ
⟧
*
Δ
vs
.
Proof
.
apply
and_proper
;
[
apply
pure_proper
;
by
rewrite
fmap_length
|
].
revert
Δ
vs
τ
i
;
induction
Γ
=>
Δ
[
|
v
vs
]
τ
i
;
csimpl
;
auto
.
apply
and_proper
;
auto
.
apply
(
interp_weaken
[]
[
τ
i
]
Δ
).
Qed
.
End
logrel
.
Typeclasses
Opaque
interp_env
.
Notation
"⟦ τ ⟧"
:=
(
interp
τ
).
Notation
"⟦ τ ⟧ₑ"
:=
(
interp_expr
τ
).
Notation
"⟦ Γ ⟧*"
:=
(
interp_env
Γ
).
F_mu_ref/soundness.v
View file @
88c9a1e0
...
...
@@ -13,7 +13,8 @@ Section soundness.
iPvs
(
heap_alloc
with
"Hemp"
)
as
{
H
}
"[Hheap Hemp]"
;
first
solve_ndisj
.
iApply
wp_wand_l
.
iSplitR
.
{
iIntros
{
v
}
"HΦ"
.
iExists
H
.
iExact
"HΦ"
.
}
rewrite
-
(
empty_env_subst
e
).
iApply
typed_interp
;
eauto
.
rewrite
-
(
empty_env_subst
e
).
iApply
fundamental
;
repeat
iSplit
;
eauto
.
by
iApply
interp_env_nil
.
Qed
.
Theorem
soundness
e
τ
e
'
thp
σ
σ'
:
...
...
F_mu_ref/typing.v
View file @
88c9a1e0
...
...
@@ -35,7 +35,7 @@ Inductive typed (Γ : list type) : expr → type → Prop :=
|
TLam_typed
e
τ
:
subst
(
ren
(
+
1
))
<
$
>
Γ
⊢ₜ
e
:
τ
→
Γ
⊢ₜ
TLam
e
:
TForall
τ
|
TApp_typed
e
τ
τ'
:
Γ
⊢ₜ
e
:
TForall
τ
→
Γ
⊢ₜ
TApp
e
:
τ
.[
τ'
/
]
|
TFold
e
τ
:
subst
(
ren
(
+
1
))
<
$
>
Γ
⊢ₜ
e
:
τ
→
Γ
⊢ₜ
Fold
e
:
TRec
τ
|
TFold
e
τ
:
Γ
⊢ₜ
e
:
τ
.[
TRec
τ
/
]
→
Γ
⊢ₜ
Fold
e
:
TRec
τ
|
TUnfold
e
τ
:
Γ
⊢ₜ
e
:
TRec
τ
→
Γ
⊢ₜ
Unfold
e
:
τ
.[
TRec
τ
/
]
|
TAlloc
e
τ
:
Γ
⊢ₜ
e
:
τ
→
Γ
⊢ₜ
Alloc
e
:
Tref
τ
|
TLoad
e
τ
:
Γ
⊢ₜ
e
:
Tref
τ
→
Γ
⊢ₜ
Load
e
:
τ
...
...
F_mu_ref_par/context_refinement.v
View file @
88c9a1e0
...
...
@@ -201,8 +201,8 @@ Section bin_log_related_under_typed_ctx.
(
∀
f
,
e
.[
base
.
iter
(
length
Γ
)
up
f
]
=
e
)
→
(
∀
f
,
e
'
.[
base
.
iter
(
length
Γ
)
up
f
]
=
e
'
)
→
typed_ctx
K
Γ
τ
Γ'
τ'
→
(
∀
Δ
(
H
Δ
:
ctx
_PersistentP
Δ
),
Δ
∥
Γ
⊨
e
≤
log
≤
e
'
:
τ
)
→
∀
Δ
(
H
Δ
:
ctx
_PersistentP
Δ
),
(
∀
Δ
(
H
Δ
:
env
_PersistentP
Δ
),
Δ
∥
Γ
⊨
e
≤
log
≤
e
'
:
τ
)
→
∀
Δ
(
H
Δ
:
env
_PersistentP
Δ
),
Δ
∥
Γ'
⊨
fill_ctx
K
e
≤
log
≤
fill_ctx
K
e
'
:
τ'
.
Proof
.
revert
Γ
τ
Γ'
τ'
e
e
'
.
...
...
@@ -211,31 +211,31 @@ 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
typed_binary_interp
_Lam
;
eauto
;
+
eapply
bin_log_related
_Lam
;
eauto
;
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_ctx_n_closed
_
_
_
_
_
_
_
H
)
end
.
+
eapply
typed_binary_interp
_App
;
eauto
using
typed_
binary_
interp
.
+
eapply
typed_binary_interp
_App
;
eauto
using
typed_
binary_
interp
.
+
eapply
typed_binary_interp
_Pair
;
eauto
using
typed_
binary_
interp
.
+
eapply
typed_binary_interp
_Pair
;
eauto
using
typed_
binary_
interp
.
+
eapply
typed_binary_interp
_Fst
;
eauto
.
+
eapply
typed_binary_interp
_Snd
;
eauto
.
+
eapply
typed_binary_interp
_InjL
;
eauto
.
+
eapply
typed_binary_interp
_InjR
;
eauto
.
+
eapply
bin_log_related
_App
;
eauto
using
binary_
fundamental
.
+
eapply
bin_log_related
_App
;
eauto
using
binary_
fundamental
.
+
eapply
bin_log_related
_Pair
;
eauto
using
binary_
fundamental
.
+
eapply
bin_log_related
_Pair
;
eauto
using
binary_
fundamental
.
+
eapply
bin_log_related
_Fst
;
eauto
.
+
eapply
bin_log_related
_Snd
;
eauto
.
+
eapply
bin_log_related
_InjL
;
eauto
.
+
eapply
bin_log_related
_InjR
;
eauto
.
+
match
goal
with
H
:
typed_ctx_item
_
_
_
_
_
|-
_
=>
inversion
H
;
subst
end
.
eapply
typed_binary_interp
_Case
;
eauto
using
typed_
binary_
interp
;
eapply
bin_log_related
_Case
;
eauto
using
binary_
fundamental
;
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_n_closed
_
_
_
H
)
end
.
+
match
goal
with
H
:
typed_ctx_item
_
_
_
_
_
|-
_
=>
inversion
H
;
subst
end
.
eapply
typed_binary_interp
_Case
;
eauto
using
typed_
binary_
interp
;
eapply
bin_log_related
_Case
;
eauto
using
binary_
fundamental
;
try
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_n_closed
_
_
_
H
)
end
;
...
...
@@ -245,36 +245,33 @@ Section bin_log_related_under_typed_ctx.
+
match
goal
with
H
:
typed_ctx_item
_
_
_
_
_
|-
_
=>
inversion
H
;
subst
end
.
eapply
typed_binary_interp
_Case
;
eauto
using
typed_
binary_
interp
;
eapply
bin_log_related
_Case
;
eauto
using
binary_
fundamental
;
try
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_n_closed
_
_
_
H
)
end
;
match
goal
with
H
:
_
|-
_
=>
eapply
(
typed_ctx_n_closed
_
_
_
_
_
_
_
H
)
end
.
+
eapply
typed_binary_interp_If
;
eauto
using
typed_ctx_typed
,
typed_binary_interp
.
+
eapply
typed_binary_interp_If
;
eauto
using
typed_ctx_typed
,
typed_binary_interp
.
+
eapply
typed_binary_interp_If
;
eauto
using
typed_ctx_typed
,
typed_binary_interp
.
+
eapply
typed_binary_interp_nat_bin_op
;
eauto
using
typed_ctx_typed
,
typed_binary_interp
.
+
eapply
typed_binary_interp_nat_bin_op
;
eauto
using
typed_ctx_typed
,
typed_binary_interp
.
+
eapply
typed_binary_interp_Fold
;
eauto
.
+
eapply
typed_binary_interp_Unfold
;
eauto
.
+
eapply
typed_binary_interp_TLam
;
eauto
with
typeclass_instances
.
+
eapply
typed_binary_interp_TApp
;
eauto
.
+
eapply
typed_binary_interp_Fork
;
eauto
.
+
eapply
typed_binary_interp_Alloc
;
eauto
.
+
eapply
typed_binary_interp_Load
;
eauto
.
+
eapply
typed_binary_interp_Store
;
eauto
using
typed_binary_interp
.
+
eapply
typed_binary_interp_Store
;
eauto
using
typed_binary_interp
.
+
eapply
typed_binary_interp_CAS
;
eauto
using
typed_binary_interp
.
+
eapply
typed_binary_interp_CAS
;
eauto
using
typed_binary_interp
.
+
eapply
typed_binary_interp_CAS
;
eauto
using
typed_binary_interp
.
+
eapply
bin_log_related_If
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_If
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_If
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_nat_bin_op
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_nat_bin_op
;
eauto
using
typed_ctx_typed
,
binary_fundamental
.
+
eapply
bin_log_related_Fold
;
eauto
.
+
eapply
bin_log_related_Unfold
;
eauto
.
+
eapply
bin_log_related_TLam
;
eauto
with
typeclass_instances
.
+
eapply
bin_log_related_TApp
;
eauto
.
+
eapply
bin_log_related_Fork
;
eauto
.
+
eapply
bin_log_related_Alloc
;
eauto
.
+
eapply
bin_log_related_Load
;
eauto
.
+
eapply
bin_log_related_Store
;
eauto
using
binary_fundamental
.