Skip to content
GitLab
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
67d0a0ab
Commit
67d0a0ab
authored
Aug 14, 2017
by
Dan Frumin
Browse files
Introduce a single typeclass for logical state for logical relations
parent
c61b7c2a
Changes
16
Hide whitespace changes
Inline
Side-by-side
F_mu_ref_conc/adequacy.v
0 → 100644
View file @
67d0a0ab
From
iris
.
program_logic
Require
Export
weakestpre
adequacy
.
From
iris_logrel
.
F_mu_ref_conc
Require
Export
rules
.
From
iris
.
algebra
Require
Import
auth
.
From
iris
.
proofmode
Require
Import
tactics
.
Set
Default
Proof
Using
"Type"
.
Class
heapPreG
Σ
:=
HeapPreG
{
heap_preG_iris
:>
invPreG
Σ
;
heap_preG_heap
:>
gen_heapPreG
loc
val
Σ
}
.
Definition
heap
Σ
:
gFunctors
:=
#[
inv
Σ
;
gen_heap
Σ
loc
val
].
Instance
subG_heapPreG
{
Σ
}
:
subG
heap
Σ
Σ
→
heapPreG
Σ
.
Proof
.
solve_inG
.
Qed
.
Definition
heap_adequacy
Σ
`
{
heapPreG
Σ
}
e
σ
φ
:
(
∀
`
{
heapG
Σ
}
,
True
⊢
WP
e
{{
v
,
⌜φ
v
⌝
}}
)
→
adequate
e
σ
φ
.
Proof
.
intros
Hwp
;
eapply
(
wp_adequacy
_
_
);
iIntros
(
?
)
""
.
iMod
(
own_alloc
(
●
to_gen_heap
σ
))
as
(
γ
)
"Hh"
.
{
apply
:
auth_auth_valid
.
exact
:
to_gen_heap_valid
.
}
iModIntro
.
iExists
(
λ
σ
,
own
γ
(
●
to_gen_heap
σ
));
iFrame
.
set
(
Hheap
:=
GenHeapG
loc
val
Σ
_
_
_
γ
).
iApply
(
Hwp
(
HeapG
_
_
_
)).
Qed
.
F_mu_ref_conc/context_refinement.v
View file @
67d0a0ab
...
...
@@ -228,7 +228,7 @@ Ltac fold_interp :=
end
.
Section
bin_log_related_under_typed_ctx
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Ltac
fundamental
:=
try
(
solve_ndisj
);
...
...
F_mu_ref_conc/examples/counter.v
View file @
67d0a0ab
...
...
@@ -29,7 +29,7 @@ Definition FG_counter : expr :=
(
FG_increment
"x"
,
counter_read
"x"
).
Section
CG_Counter
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
(
*
Coarse
-
grained
increment
*
)
Lemma
CG_increment_type
Γ
:
...
...
@@ -314,12 +314,6 @@ Theorem counter_ctx_refinement :
∅
⊨
FG_counter
≤
ctx
≤
CG_counter
:
TProd
(
TArrow
TUnit
TUnit
)
(
TArrow
TUnit
TNat
).
Proof
.
set
(
Σ
:=
#[
inv
Σ
;
gen_heap
Σ
loc
val
;
auth
Σ
cfgUR
]).
set
(
HG
:=
HeapPreIG
Σ
_
_
).
eapply
(
logrel_ctxequiv
Σ
_
).
(
*
TODO
:
how
to
get
rid
of
this
bullshit
with
closed
conditions
?
*
)
rewrite
/
FG_counter
/
CG_counter
;
try
solve_closed
.
rewrite
/
FG_counter
/
CG_counter
;
try
solve_closed
.
Transparent
newlock
.
unfold
newlock
.
solve_closed
.
intros
.
apply
FG_CG_counter_refinement
.
eapply
(
logrel_ctxequiv
logrel
Σ
);
[
solve_closed
..
|
intros
].
apply
FG_CG_counter_refinement
.
Qed
.
F_mu_ref_conc/examples/lateearlychoice.v
View file @
67d0a0ab
...
...
@@ -17,7 +17,7 @@ Definition earlyChoice : val := λ: "x",
let:
"r"
:=
rand
#()
in
"x"
<-
#
n
0
;;
"r"
.
Section
Refinement
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Definition
choiceN
:
namespace
:=
nroot
.
@
"choice"
.
...
...
F_mu_ref_conc/examples/lock.v
View file @
67d0a0ab
...
...
@@ -52,8 +52,7 @@ Qed.
Hint
Resolve
with_lock_type
:
typeable
.
Section
proof
.
Context
`
{
cfgSG
Σ
}
.
Context
`
{
heapIG
Σ
}
.
Context
`
{
logrelG
Σ
}
.
Variable
(
E1
E2
:
coPset
).
Lemma
steps_newlock
ρ
j
K
...
...
F_mu_ref_conc/examples/par.v
View file @
67d0a0ab
...
...
@@ -29,7 +29,7 @@ Qed.
Hint
Resolve
par_type
:
typeable
.
Section
compatibility
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Lemma
bin_log_related_par
Γ
E
e1
e2
e1
'
e2
'
τ
1
τ
2
:
↑
specN
⊆
E
→
...
...
F_mu_ref_conc/fundamental_binary.v
View file @
67d0a0ab
...
...
@@ -5,7 +5,7 @@ From iris.base_logic Require Export big_op.
From
iris
.
program_logic
Require
Import
ectx_lifting
.
Section
fundamental
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Notation
D
:=
(
prodC
valC
valC
-
n
>
iProp
Σ
).
Implicit
Types
e
:
expr
.
Implicit
Types
Δ
:
listC
D
.
...
...
F_mu_ref_conc/hax.v
View file @
67d0a0ab
...
...
@@ -25,7 +25,7 @@ Ltac inv_head_step :=
end
.
Section
hax
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Notation
D
:=
(
prodC
valC
valC
-
n
>
iProp
Σ
).
Implicit
Types
Δ
:
listC
D
.
...
...
F_mu_ref_conc/logrel_binary.v
View file @
67d0a0ab
...
...
@@ -38,7 +38,7 @@ Definition logN : namespace := nroot .@ "logN".
(
**
interp
:
is
a
unary
logical
relation
.
*
)
Section
logrel
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Notation
D
:=
(
prodC
valC
valC
-
n
>
iProp
Σ
).
Implicit
Types
τ
i
:
D
.
Implicit
Types
Δ
:
listC
D
.
...
...
@@ -372,7 +372,7 @@ Notation "⟦ τ ⟧ₑ" := (interp_expr ⊤ ⊤ ⟦ τ ⟧).
Notation
"⟦ Γ ⟧*"
:=
(
interp_env
Γ
).
Section
bin_log_def
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Notation
D
:=
(
prodC
valC
valC
-
n
>
iProp
Σ
).
Definition
bin_log_related_def
(
E1
E2
:
coPset
)
(
Γ
:
stringmap
type
)
(
e
e
'
:
expr
)
(
τ
:
type
)
:
iProp
Σ
:=
(
∀
Δ
(
vvs
:
stringmap
(
val
*
val
))
ρ
,
...
...
F_mu_ref_conc/rel_tactics.v
View file @
67d0a0ab
...
...
@@ -10,7 +10,7 @@ Import lang.
LHS
once
if
necessary
,
to
get
rid
of
the
lock
added
by
the
syntactic
sugar
.
*
)
Ltac
solve_of_val_unlock
:=
try
apply
of_val_unlock
;
fast_done
.
Lemma
tac_rel_bind_gen
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
e
'
t
t
'
τ
:
Lemma
tac_rel_bind_gen
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
e
'
t
t
'
τ
:
e
=
e
'
→
t
=
t
'
→
(
Δ
⊢
bin_log_related
E1
E2
Γ
e
'
t
'
τ
)
→
...
...
@@ -19,13 +19,13 @@ Proof.
intros
.
subst
t
e
.
assumption
.
Qed
.
Lemma
tac_rel_bind_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
e
'
K
Δ
E1
E2
Γ
e
t
τ
:
Lemma
tac_rel_bind_l
`
{
logrel
G
Σ
}
e
'
K
Δ
E1
E2
Γ
e
t
τ
:
e
=
fill
K
e
'
→
(
Δ
⊢
bin_log_related
E1
E2
Γ
(
fill
K
e
'
)
t
τ
)
→
(
Δ
⊢
bin_log_related
E1
E2
Γ
e
t
τ
).
Proof
.
intros
.
eapply
tac_rel_bind_gen
;
eauto
.
Qed
.
Lemma
tac_rel_bind_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
t
'
K
Δ
E1
E2
Γ
e
t
τ
:
Lemma
tac_rel_bind_r
`
{
logrel
G
Σ
}
t
'
K
Δ
E1
E2
Γ
e
t
τ
:
t
=
fill
K
t
'
→
(
Δ
⊢
bin_log_related
E1
E2
Γ
e
(
fill
K
t
'
)
τ
)
→
(
Δ
⊢
bin_log_related
E1
E2
Γ
e
t
τ
).
...
...
@@ -76,7 +76,7 @@ Tactic Notation "rel_bind_r" open_constr(efoc) :=
|
(
*
new
goal
*
)
].
Lemma
tac_rel_rec_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
K
'
f
x
ef
e
'
efbody
v
eres
t
τ
:
Lemma
tac_rel_rec_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
K
'
f
x
ef
e
'
efbody
v
eres
t
τ
:
e
=
fill
K
'
(
App
ef
e
'
)
→
ef
=
Rec
f
x
efbody
→
Closed
(
x
:
b
:
f
:
b
:
∅
)
efbody
→
...
...
@@ -110,7 +110,7 @@ Tactic Notation "rel_rec_l" :=
Tactic
Notation
"rel_seq_l"
:=
rel_rec_l
.
Tactic
Notation
"rel_let_l"
:=
rel_rec_l
.
Lemma
tac_rel_fst_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e1
e2
v1
v2
K
'
t
τ
:
Lemma
tac_rel_fst_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e1
e2
v1
v2
K
'
t
τ
:
e
=
fill
K
'
(
Fst
(
Pair
e1
e2
))
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
...
...
@@ -132,7 +132,7 @@ Tactic Notation "rel_fst_l" :=
|
solve_to_val
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_snd_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e1
e2
v1
v2
K
'
t
τ
:
Lemma
tac_rel_snd_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e1
e2
v1
v2
K
'
t
τ
:
e
=
fill
K
'
(
Snd
(
Pair
e1
e2
))
→
to_val
e1
=
Some
v1
→
to_val
e2
=
Some
v2
→
...
...
@@ -154,7 +154,7 @@ Tactic Notation "rel_snd_l" :=
|
solve_to_val
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_unfold_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e1
v
K
'
t
τ
:
Lemma
tac_rel_unfold_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e1
v
K
'
t
τ
:
e
=
fill
K
'
(
Unfold
(
Fold
e1
))
→
to_val
e1
=
Some
v
→
(
Δ
⊢
▷
bin_log_related
E1
E1
Γ
(
fill
K
'
e1
)
t
τ
)
→
...
...
@@ -174,7 +174,7 @@ Tactic Notation "rel_unfold_l" :=
Tactic
Notation
"rel_fold_l"
:=
rel_unfold_l
.
Lemma
tac_rel_if_true_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e1
e2
K
'
t
τ
:
Lemma
tac_rel_if_true_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e1
e2
K
'
t
τ
:
e
=
fill
K
'
(
If
(#
♭
true
)
e1
e2
)
→
Closed
∅
e1
→
Closed
∅
e2
→
...
...
@@ -194,7 +194,7 @@ Tactic Notation "rel_if_true_l" :=
|
try
solve_closed
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_if_false_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e1
e2
K
'
t
τ
:
Lemma
tac_rel_if_false_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e1
e2
K
'
t
τ
:
e
=
fill
K
'
(
If
(#
♭
false
)
e1
e2
)
→
Closed
∅
e1
→
Closed
∅
e2
→
...
...
@@ -214,7 +214,7 @@ Tactic Notation "rel_if_false_l" :=
|
try
solve_closed
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_case_inl_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e0
v
e1
e2
K
'
t
τ
:
Lemma
tac_rel_case_inl_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e0
v
e1
e2
K
'
t
τ
:
e
=
fill
K
'
(
Case
(
InjL
e0
)
e1
e2
)
→
to_val
e0
=
Some
v
→
Closed
∅
e1
→
...
...
@@ -236,7 +236,7 @@ Tactic Notation "rel_case_inl_l" :=
|
try
solve_closed
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_case_inr_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
e0
v
e1
e2
K
'
t
τ
:
Lemma
tac_rel_case_inr_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
e0
v
e1
e2
K
'
t
τ
:
e
=
fill
K
'
(
Case
(
InjR
e0
)
e1
e2
)
→
to_val
e0
=
Some
v
→
Closed
∅
e1
→
...
...
@@ -258,7 +258,7 @@ Tactic Notation "rel_case_inr_l" :=
|
try
solve_closed
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_binop_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
Γ
e
K
'
op
a
b
eres
t
τ
:
Lemma
tac_rel_binop_l
`
{
logrel
G
Σ
}
Δ
E1
Γ
e
K
'
op
a
b
eres
t
τ
:
e
=
fill
K
'
(
BinOp
op
(#
n
a
)
(#
n
b
))
→
eres
=
of_val
(
binop_eval
op
a
b
)
→
(
Δ
⊢
▷
bin_log_related
E1
E1
Γ
(
fill
K
'
eres
)
t
τ
)
→
...
...
@@ -276,7 +276,7 @@ Tactic Notation "rel_op_l" :=
|
simpl
;
reflexivity
(
*
eres
=
of_val
..
*
)
|
iNext
(
*
new
goal
*
)].
Lemma
tac_rel_fork_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
E1
E2
e
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_fork_l
`
{
logrel
G
Σ
}
Δ
1
E1
E2
e
'
K
'
Γ
e
t
τ
:
e
=
fill
K
'
(
Fork
e
'
)
→
Closed
∅
e
'
→
(
Δ
1
⊢
|={
E1
,
E2
}=>
▷
WP
e
'
{{
_
,
True
}}
∗
bin_log_related
E2
E1
Γ
(
fill
K
'
(
Lit
Unit
))
t
τ
)
→
...
...
@@ -294,7 +294,7 @@ Tactic Notation "rel_fork_l" :=
|
solve_closed
|
simpl
(
*
new
goal
*
)
].
Lemma
tac_rel_alloc_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
e
'
v
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_alloc_l
`
{
logrel
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
e
'
v
'
K
'
Γ
e
t
τ
:
nclose
N
⊆
E1
→
envs_lookup
i1
Δ
1
=
Some
(
p
,
inv
N
P
)
→
E2
=
E1
∖
↑
N
→
...
...
@@ -345,7 +345,7 @@ Tactic Notation "rel_alloc_l" "under" constr(N) "as" constr(nam) constr(nam_cl)
|
env_cbv
;
reflexivity
||
fail
"rel_alloc_l: this should not happen"
|
(
*
new
goal
*
)].
Lemma
tac_rel_alloc_l_simp
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
E1
e
'
v
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_alloc_l_simp
`
{
logrel
G
Σ
}
Δ
1
Δ
2
E1
e
'
v
'
K
'
Γ
e
t
τ
:
e
=
fill
K
'
(
Alloc
e
'
)
→
to_val
e
'
=
Some
v
'
→
(
Δ
1
⊢
∀
l
,
...
...
@@ -366,7 +366,7 @@ Tactic Notation "rel_alloc_l" "as" ident(l) constr(H) :=
|
iIntros
(
l
)
H
(
*
new
goal
*
)].
Lemma
tac_rel_load_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
l
K
'
Γ
e
t
τ
:
Lemma
tac_rel_load_l
`
{
logrel
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
l
K
'
Γ
e
t
τ
:
nclose
N
⊆
E1
→
envs_lookup
i1
Δ
1
=
Some
(
p
,
inv
N
P
)
→
E2
=
E1
∖
↑
N
→
...
...
@@ -416,7 +416,7 @@ Tactic Notation "rel_load_l" "under" constr(N) "as" constr(nam) constr(nam_cl) :
|
env_cbv
;
reflexivity
||
fail
"rel_load_l: this should not happen"
|
(
*
new
goal
*
)].
Lemma
tac_rel_load_l_simp
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
Δ
2
E1
i1
l
v
K
'
Γ
e
t
τ
:
Lemma
tac_rel_load_l_simp
`
{
logrel
G
Σ
}
Δ
1
Δ
2
E1
i1
l
v
K
'
Γ
e
t
τ
:
e
=
fill
K
'
(
Load
(
Loc
l
))
→
IntoLaterNEnvs
1
Δ
1
Δ
2
→
envs_lookup
i1
Δ
2
=
Some
(
false
,
l
↦ᵢ
v
)
%
I
→
...
...
@@ -440,7 +440,7 @@ Tactic Notation "rel_load_l" :=
|
iAssumptionCore
||
fail
3
"rel_load_l: cannot find ? ↦ᵢ ?"
|
(
*
new
goal
*
)].
Lemma
tac_rel_store_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
l
e
'
v
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_store_l
`
{
logrel
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
l
e
'
v
'
K
'
Γ
e
t
τ
:
nclose
N
⊆
E1
→
envs_lookup
i1
Δ
1
=
Some
(
p
,
inv
N
P
)
→
E2
=
E1
∖
↑
N
→
...
...
@@ -493,7 +493,7 @@ Tactic Notation "rel_store_l" "under" constr(N) "as" constr(nam) constr(nam_cl)
|
(
*
new
goal
*
)].
Lemma
tac_rel_store_l_simp
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
Δ
2
i1
E1
l
v
e
'
v
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_store_l_simp
`
{
logrel
G
Σ
}
Δ
1
Δ
2
i1
E1
l
v
e
'
v
'
K
'
Γ
e
t
τ
:
e
=
fill
K
'
(
Store
(
Loc
l
)
e
'
)
→
to_val
e
'
=
Some
v
'
→
envs_lookup
i1
Δ
1
=
Some
(
false
,
l
↦ᵢ
v
)
%
I
→
...
...
@@ -520,7 +520,7 @@ Tactic Notation "rel_store_l" :=
|
env_cbv
;
reflexivity
||
fail
"rel_store_l: this should not happen"
|
(
*
new
goal
*
)].
Lemma
tac_rel_cas_l
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
l
e1
e2
v1
v2
K
'
Γ
e
t
τ
:
Lemma
tac_rel_cas_l
`
{
logrel
G
Σ
}
nam
nam_cl
Δ
1
Δ
2
E1
E2
p
i1
N
P
l
e1
e2
v1
v2
K
'
Γ
e
t
τ
:
nclose
N
⊆
E1
→
envs_lookup
i1
Δ
1
=
Some
(
p
,
inv
N
P
)
→
E2
=
E1
∖
↑
N
→
...
...
@@ -584,7 +584,7 @@ Tactic Notation "rel_cas_l" "under" constr(N) "as" constr(nam) constr(nam_cl) :=
(
********************************
)
Lemma
tac_rel_fork_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
E1
E2
t
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_fork_r
`
{
logrel
G
Σ
}
Δ
1
E1
E2
t
'
K
'
Γ
e
t
τ
:
nclose
specN
⊆
E1
→
t
=
fill
K
'
(
Fork
t
'
)
→
Closed
∅
t
'
→
...
...
@@ -604,7 +604,7 @@ Tactic Notation "rel_fork_r" "as" ident(i) constr(H) :=
|
solve_closed
|
simpl
;
iIntros
(
i
)
H
(
*
new
goal
*
)].
Lemma
tac_rel_store_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
t
'
v
'
K
'
Γ
e
t
τ
v
:
Lemma
tac_rel_store_r
`
{
logrel
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
t
'
v
'
K
'
Γ
e
t
τ
v
:
nclose
specN
⊆
E1
→
t
=
fill
K
'
(
Store
(
Loc
l
)
t
'
)
→
to_val
t
'
=
Some
v
'
→
...
...
@@ -632,7 +632,7 @@ Tactic Notation "rel_store_r" :=
|
env_cbv
;
reflexivity
||
fail
"rel_store_r: this should not happen"
|
(
*
new
goal
*
)].
Lemma
tac_rel_alloc_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
E1
E2
t
'
v
'
K
'
Γ
e
t
τ
:
Lemma
tac_rel_alloc_r
`
{
logrel
G
Σ
}
Δ
1
E1
E2
t
'
v
'
K
'
Γ
e
t
τ
:
nclose
specN
⊆
E1
→
t
=
fill
K
'
(
Alloc
t
'
)
→
to_val
t
'
=
Some
v
'
→
...
...
@@ -657,7 +657,7 @@ Tactic Notation "rel_alloc_r" :=
let
H
:=
iFresh
"H"
in
rel_alloc_r
as
l
H
.
Lemma
tac_rel_load_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
K
'
Γ
e
t
τ
v
:
Lemma
tac_rel_load_r
`
{
logrel
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
K
'
Γ
e
t
τ
v
:
nclose
specN
⊆
E1
→
t
=
fill
K
'
(
Load
(
Loc
l
))
→
envs_lookup
i1
Δ
1
=
Some
(
false
,
l
↦ₛ
v
)
%
I
→
...
...
@@ -684,7 +684,7 @@ Tactic Notation "rel_load_r" :=
|
env_cbv
;
reflexivity
||
fail
"rel_load_r: this should not happen"
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_cas_fail_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
K
'
Γ
e
t
e1
e2
v1
v2
τ
v
:
Lemma
tac_rel_cas_fail_r
`
{
logrel
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
K
'
Γ
e
t
e1
e2
v1
v2
τ
v
:
nclose
specN
⊆
E1
→
t
=
fill
K
'
(
CAS
(
Loc
l
)
e1
e2
)
→
to_val
e1
=
Some
v1
→
...
...
@@ -718,7 +718,7 @@ Tactic Notation "rel_cas_fail_r" :=
|
(
*
new
goal
*
)].
Lemma
tac_rel_cas_suc_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
K
'
Γ
e
t
e1
e2
v1
v2
τ
v
:
Lemma
tac_rel_cas_suc_r
`
{
logrel
G
Σ
}
Δ
1
Δ
2
E1
E2
i1
l
K
'
Γ
e
t
e1
e2
v1
v2
τ
v
:
nclose
specN
⊆
E1
→
t
=
fill
K
'
(
CAS
(
Loc
l
)
e1
e2
)
→
to_val
e1
=
Some
v1
→
...
...
@@ -752,7 +752,7 @@ Tactic Notation "rel_cas_suc_r" :=
|
(
*
new
goal
*
)].
Lemma
tac_rel_rec_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
f
x
ef
e
'
efbody
v
eres
t
τ
:
Lemma
tac_rel_rec_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
f
x
ef
e
'
efbody
v
eres
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
App
ef
e
'
)
→
ef
=
Rec
f
x
efbody
→
...
...
@@ -788,7 +788,7 @@ Tactic Notation "rel_rec_r" :=
Tactic
Notation
"rel_seq_r"
:=
rel_rec_r
.
Tactic
Notation
"rel_let_r"
:=
rel_rec_r
.
Lemma
tac_rel_fst_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
v1
v2
t
τ
:
Lemma
tac_rel_fst_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
v1
v2
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
Fst
(
Pair
e1
e2
))
→
to_val
e1
=
Some
v1
→
...
...
@@ -813,7 +813,7 @@ Tactic Notation "rel_fst_r" :=
|
solve_to_val
(
*
to_val
e2
=
Some
..
*
)
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_snd_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
v1
v2
t
τ
:
Lemma
tac_rel_snd_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
v1
v2
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
Snd
(
Pair
e1
e2
))
→
to_val
e1
=
Some
v1
→
...
...
@@ -838,7 +838,7 @@ Tactic Notation "rel_snd_r" :=
|
solve_to_val
(
*
to_val
e2
=
Some
..
*
)
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_tlam_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e
'
t
τ
:
Lemma
tac_rel_tlam_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e
'
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
TApp
(
TLam
e
'
))
→
Closed
∅
e
'
→
...
...
@@ -858,7 +858,7 @@ Tactic Notation "rel_tlam_r" :=
|
solve_closed
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_fold_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e
'
v
t
τ
:
Lemma
tac_rel_fold_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e
'
v
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
Unfold
(
Fold
e
'
))
→
to_val
e
'
=
Some
v
→
...
...
@@ -880,7 +880,7 @@ Tactic Notation "rel_fold_r" :=
Tactic
Notation
"rel_unfold_r"
:=
rel_fold_r
.
Lemma
tac_rel_case_inl_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e0
e1
e2
v
t
τ
:
Lemma
tac_rel_case_inl_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e0
e1
e2
v
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
Case
(
InjL
e0
)
e1
e2
)
→
Closed
∅
e1
→
...
...
@@ -904,7 +904,7 @@ Tactic Notation "rel_case_inl_r" :=
|
solve_to_val
(
*
to_val
e0
=
Some
..
*
)
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_case_inr_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e0
e1
e2
v
t
τ
:
Lemma
tac_rel_case_inr_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e0
e1
e2
v
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
Case
(
InjR
e0
)
e1
e2
)
→
Closed
∅
e1
→
...
...
@@ -930,7 +930,7 @@ Tactic Notation "rel_case_inr_r" :=
Tactic
Notation
"rel_case_r"
:=
rel_case_inl_r
||
rel_case_inr_r
.
Lemma
tac_rel_if_true_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
t
τ
:
Lemma
tac_rel_if_true_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
If
(#
♭
true
)
e1
e2
)
→
Closed
∅
e1
→
...
...
@@ -952,7 +952,7 @@ Tactic Notation "rel_if_true_r" :=
|
solve_closed
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_if_false_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
t
τ
:
Lemma
tac_rel_if_false_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
e1
e2
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
If
(#
♭
false
)
e1
e2
)
→
Closed
∅
e1
→
...
...
@@ -974,7 +974,7 @@ Tactic Notation "rel_if_false_r" :=
|
solve_closed
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_if_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
b
eres
e1
e2
t
τ
:
Lemma
tac_rel_if_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
b
eres
e1
e2
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
If
(#
♭
b
)
e1
e2
)
→
Closed
∅
e1
→
...
...
@@ -1000,7 +1000,7 @@ Tactic Notation "rel_if_r" :=
|
simpl
;
fast_done
||
fail
"rel_if_r: cannot compute the boolean value"
|
simpl
(
*
new
goal
*
)].
Lemma
tac_rel_binop_r
`
{
heapIG
Σ
,
!
cfgS
G
Σ
}
Δ
E1
E2
Γ
e
K
'
op
a
b
t
τ
:
Lemma
tac_rel_binop_r
`
{
logrel
G
Σ
}
Δ
E1
E2
Γ
e
K
'
op
a
b
t
τ
:
nclose
specN
⊆
E1
→
e
=
fill
K
'
(
BinOp
op
(#
n
a
)
(#
n
b
))
→
(
Δ
⊢
bin_log_related
E1
E2
Γ
t
(
fill
K
'
(
of_val
(
binop_eval
op
a
b
)))
τ
)
→
...
...
@@ -1027,7 +1027,7 @@ Tactic Notation "rel_vals" :=
iIntros
(
d
);
simpl
.
Section
test
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Definition
choiceN
:
namespace
:=
nroot
.
@
"choice"
.
...
...
F_mu_ref_conc/relational_properties.v
View file @
67d0a0ab
...
...
@@ -5,7 +5,7 @@ From iris.program_logic Require Import ectx_lifting.
(
**
*
Properties
of
the
relational
interpretation
*
)
Section
properties
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Notation
D
:=
(
prodC
valC
valC
-
n
>
iProp
Σ
).
Implicit
Types
e
:
expr
.
Implicit
Types
Δ
:
listC
D
.
...
...
F_mu_ref_conc/rules.v
View file @
67d0a0ab
...
...
@@ -9,13 +9,13 @@ Import uPred.
(
**
The
CMRA
for
the
heap
of
the
implementation
.
This
is
linked
to
the
physical
heap
.
*
)
Class
heap
I
G
Σ
:=
Heap
I
G
{
heap
I
_invG
:
invG
Σ
;
heap
I
_gen_heapG
:>
gen_heapG
loc
val
Σ
;
Class
heapG
Σ
:=
HeapG
{
heap
G
_invG
:
invG
Σ
;
heap
G
_gen_heapG
:>
gen_heapG
loc
val
Σ
;
}
.
Instance
heap
I
G_irisG
`
{
heap
I
G
Σ
}
:
irisG
lang
Σ
:=
{
iris_invG
:=
heap
I
_invG
;
Instance
heapG_irisG
`
{
heapG
Σ
}
:
irisG
lang
Σ
:=
{
iris_invG
:=
heap
G
_invG
;
state_interp
:=
gen_heap_ctx
}
.
Global
Opaque
iris_invG
.
...
...
@@ -25,7 +25,7 @@ Notation "l ↦ᵢ{ q } v" := (mapsto (L:=loc) (V:=val) l q v)
Notation
"l ↦ᵢ v"
:=
(
mapsto
(
L
:=
loc
)
(
V
:=
val
)
l
1
v
)
(
at
level
20
)
:
uPred_scope
.
Section
lang_rules
.
Context
`
{
heap
I
G
Σ
}
.
Context
`
{
heapG
Σ
}
.
Implicit
Types
P
Q
:
iProp
Σ
.
Implicit
Types
Φ
:
val
→
iProp
Σ
.
Implicit
Types
σ
:
state
.
...
...
F_mu_ref_conc/rules_binary.v
View file @
67d0a0ab
...
...
@@ -19,10 +19,14 @@ Fixpoint to_tpool_go (i : nat) (tp : list expr) : tpoolUR :=
Definition
to_tpool
:
list
expr
→
tpoolUR
:=
to_tpool_go
0.
(
**
The
CMRA
for
the
thread
pool
.
*
)
Class
cfgSG
Σ
:=
CFGSG
{
cfg_inG
:>
inG
Σ
(
authR
cfgUR
);
cfg_name
:
gname
}
.
Class
logrelG
Σ
:=
LogrelG
{
heap_inG
:>
heapG
Σ
;
cfg_inG
:>
inG
Σ
(
authR
cfgUR
);
cfg_name
:
gname
}
.
Section
definitionsS
.
Context
`
{
cfgSG
Σ
,
inv
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Definition
heapS_mapsto_def
(
l
:
loc
)
(
q
:
Qp
)
(
v
:
val
)
:
iProp
Σ
:=
own
cfg_name
(
◯
(
∅
,
{
[
l
:=
(
q
,
to_agree
v
)
]
}
)).
...
...
@@ -51,7 +55,7 @@ Notation "l ↦ₛ v" := (heapS_mapsto l 1 v) (at level 20) : uPred_scope.
Notation
"j ⤇ e"
:=
(
tpool_mapsto
j
e
)
(
at
level
20
)
:
uPred_scope
.
Section
conversions
.
Context
`
{
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
(
**
Conversion
to
tpools
and
back
*
)
Lemma
to_tpool_valid
es
:
✓
to_tpool
es
.
...
...
@@ -112,7 +116,7 @@ Section conversions.
End
conversions
.
Section
cfg
.
Context
`
{
heapIG
Σ
,
cfgS
G
Σ
}
.
Context
`
{
logrel
G
Σ
}
.
Implicit
Types
P
Q
:
iProp
Σ
.
Implicit
Types
Φ
:
val
→
iProp
Σ
.
Implicit
Types
σ
:
state
.
...
...
F_mu_ref_conc/soundness_binary.v
View file @
67d0a0ab
From
iris_logrel
.
F_mu_ref_conc
Require
Export
context_refinement
.
From
iris_logrel
.
F_mu_ref_conc
Require
Export
context_refinement
adequacy
.
From
iris
.
algebra
Require
Import
auth
frac
agree
.
From
iris
.
base_logic
Require
Import
big_op
.
From
iris
.
base_logic
Require
Import
big_op
lib
.
auth
.
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
program_logic
Require
Import
adequacy
.
Class
heap
Pre
I
G
Σ
:=
Heap
Pre
I
G
{
heap_preG_iris
:>
inv
PreG
Σ
;
heap_preG_heap
:>
gen_heapPreG
loc
val
Σ
Class
logrel
PreG
Σ
:=
Logrel
PreG
{
logrel_preG_heap
:>
heap
PreG
Σ
;
logrel_preG_cfg
:>
inG
Σ
(
authR
cfgUR
)
}
.
Lemma
logrel_adequate
Σ
`
{
heapPreIG
Σ
,
inG
Σ
(
authR
cfgUR
)
}
Definition
logrel
Σ
:
gFunctors
:=
#[
heap
Σ
;
auth
Σ
cfgUR
].
Instance
subG_heapPreG
{
Σ
}
:
subG
logrel
Σ
Σ
→
logrelPreG
Σ
.
Proof
.
solve_inG
.
Qed
.
Lemma
logrel_adequate
Σ
`
{
logrelPreG
Σ
}
e
e
'
τ
σ
:
(
∀
`
{
heapIG
Σ
,
cfgS
G
Σ
}
,
∅
⊨
e
≤
log
≤
e
'
:
τ
)
→
(
∀
`
{
logrel
G
Σ
}
,
∅
⊨
e
≤
log
≤
e
'
:
τ
)
→
adequate
e
σ
(
λ
v
,
∃
thp
'
h
v
'
,
rtc
step
([
e
'
],
∅
)
(
of_val
v
'
::
thp
'
,
h
)
∧
(
ObsType
τ
→
v
=
v
'
)).
Proof
.
intros
Hlog
.
eapply
(
wp_adequacy
Σ
_
);
iIntros
(
Hinv
).
iMod
(
own_alloc
(
●
to_gen_heap
σ
))
as
(
γ
)
"Hh"
.
{
apply
(
auth_auth_valid
_
(
to_gen_heap_valid
_
_
σ
)).
}
eapply
(
heap_adequacy
Σ
_
);
iIntros
(
Hinv
).
iMod
(
own_alloc
(
●
(
to_tpool
[
e
'
],
∅
)
⋅
◯
((
to_tpool
[
e
'
]
:
tpoolUR
,
∅
)
:
cfgUR
)))
as
(
γ
c
)
"[Hcfg1 Hcfg2]"
.
{
apply
auth_valid_discrete_2
.
split
=>
//. split=>//. apply to_tpool_valid. }
set
(
Hcfg
:=
CFGSG
_
_
γ
c
).
set
(
Hcfg
:=
LogrelG
_
_
_
γ
c
).
iMod
(
inv_alloc
specN
_
(
spec_inv
([
e
'
],
∅
))
with
"[Hcfg1]"
)
as
"#Hcfg"
.
{
iNext
.
iExists
[
e
'
],
∅
.
rewrite
/
to_gen_heap
fin_maps
.
map_fmap_empty
.
auto
.
}
set
(
Heap
Σ
:=
(
HeapIG
Σ
Hinv
(
GenHeapG
_
_
Σ
_
_
_
γ
))).
iExists
(
λ
σ
,
own
γ
(
●
to_gen_heap
σ
));
iFrame
.
iApply
wp_fupd
.
iApply
wp_wand_r
.
iSplitL
.