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
Rodolphe Lepigre
Iris
Commits
2d9c5f33
Commit
2d9c5f33
authored
Feb 25, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Seal off all definitions in uPred.
The performance gain seems neglectable, unfortunatelly...
parent
52d7d275
Changes
15
Expand all
Hide whitespace changes
Inline
Side-by-side
Showing
15 changed files
with
244 additions
and
142 deletions
+244
-142
algebra/agree.v
algebra/agree.v
+4
-2
algebra/auth.v
algebra/auth.v
+2
-2
algebra/excl.v
algebra/excl.v
+4
-2
algebra/fin_maps.v
algebra/fin_maps.v
+2
-2
algebra/iprod.v
algebra/iprod.v
+2
-2
algebra/option.v
algebra/option.v
+4
-2
algebra/upred.v
algebra/upred.v
+186
-100
program_logic/adequacy.v
program_logic/adequacy.v
+6
-6
program_logic/ghost_ownership.v
program_logic/ghost_ownership.v
+1
-1
program_logic/lifting.v
program_logic/lifting.v
+7
-5
program_logic/ownership.v
program_logic/ownership.v
+5
-2
program_logic/pviewshifts.v
program_logic/pviewshifts.v
+10
-9
program_logic/resources.v
program_logic/resources.v
+4
-2
program_logic/sts.v
program_logic/sts.v
+1
-1
program_logic/weakestpre.v
program_logic/weakestpre.v
+6
-4
No files found.
algebra/agree.v
View file @
2d9c5f33
...
@@ -134,9 +134,11 @@ Proof. intros [??]; split; naive_solver eauto using agree_valid_le. Qed.
...
@@ -134,9 +134,11 @@ Proof. intros [??]; split; naive_solver eauto using agree_valid_le. Qed.
(** Internalized properties *)
(** Internalized properties *)
Lemma
agree_equivI
{
M
}
a
b
:
(
to_agree
a
≡
to_agree
b
)%
I
≡
(
a
≡
b
:
uPred
M
)%
I
.
Lemma
agree_equivI
{
M
}
a
b
:
(
to_agree
a
≡
to_agree
b
)%
I
≡
(
a
≡
b
:
uPred
M
)%
I
.
Proof
.
do
2
split
.
by
intros
[?
Hv
]
;
apply
(
Hv
n
).
apply
:
to_agree_ne
.
Qed
.
Proof
.
uPred
.
unseal
.
do
2
split
.
by
intros
[?
Hv
]
;
apply
(
Hv
n
).
apply
:
to_agree_ne
.
Qed
.
Lemma
agree_validI
{
M
}
x
y
:
✓
(
x
⋅
y
)
⊑
(
x
≡
y
:
uPred
M
).
Lemma
agree_validI
{
M
}
x
y
:
✓
(
x
⋅
y
)
⊑
(
x
≡
y
:
uPred
M
).
Proof
.
split
=>
r
n
_
?
;
by
apply
:
agree_op_inv
.
Qed
.
Proof
.
uPred
.
unseal
;
split
=>
r
n
_
?
;
by
apply
:
agree_op_inv
.
Qed
.
End
agree
.
End
agree
.
Arguments
agreeC
:
clear
implicits
.
Arguments
agreeC
:
clear
implicits
.
...
...
algebra/auth.v
View file @
2d9c5f33
...
@@ -151,14 +151,14 @@ Qed.
...
@@ -151,14 +151,14 @@ Qed.
(** Internalized properties *)
(** Internalized properties *)
Lemma
auth_equivI
{
M
}
(
x
y
:
auth
A
)
:
Lemma
auth_equivI
{
M
}
(
x
y
:
auth
A
)
:
(
x
≡
y
)%
I
≡
(
authoritative
x
≡
authoritative
y
∧
own
x
≡
own
y
:
uPred
M
)%
I
.
(
x
≡
y
)%
I
≡
(
authoritative
x
≡
authoritative
y
∧
own
x
≡
own
y
:
uPred
M
)%
I
.
Proof
.
done
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
Lemma
auth_validI
{
M
}
(
x
:
auth
A
)
:
Lemma
auth_validI
{
M
}
(
x
:
auth
A
)
:
(
✓
x
)%
I
≡
(
match
authoritative
x
with
(
✓
x
)%
I
≡
(
match
authoritative
x
with
|
Excl
a
=>
(
∃
b
,
a
≡
own
x
⋅
b
)
∧
✓
a
|
Excl
a
=>
(
∃
b
,
a
≡
own
x
⋅
b
)
∧
✓
a
|
ExclUnit
=>
✓
own
x
|
ExclUnit
=>
✓
own
x
|
ExclBot
=>
False
|
ExclBot
=>
False
end
:
uPred
M
)%
I
.
end
:
uPred
M
)%
I
.
Proof
.
by
destruct
x
as
[[]].
Qed
.
Proof
.
uPred
.
unseal
.
by
destruct
x
as
[[]].
Qed
.
(** The notations ◯ and ● only work for CMRAs with an empty element. So, in
(** The notations ◯ and ● only work for CMRAs with an empty element. So, in
what follows, we assume we have an empty element. *)
what follows, we assume we have an empty element. *)
...
...
algebra/excl.v
View file @
2d9c5f33
...
@@ -146,10 +146,12 @@ Lemma excl_equivI {M} (x y : excl A) :
...
@@ -146,10 +146,12 @@ Lemma excl_equivI {M} (x y : excl A) :
|
ExclUnit
,
ExclUnit
|
ExclBot
,
ExclBot
=>
True
|
ExclUnit
,
ExclUnit
|
ExclBot
,
ExclBot
=>
True
|
_
,
_
=>
False
|
_
,
_
=>
False
end
:
uPred
M
)%
I
.
end
:
uPred
M
)%
I
.
Proof
.
do
2
split
.
by
destruct
1
.
by
destruct
x
,
y
;
try
constructor
.
Qed
.
Proof
.
uPred
.
unseal
.
do
2
split
.
by
destruct
1
.
by
destruct
x
,
y
;
try
constructor
.
Qed
.
Lemma
excl_validI
{
M
}
(
x
:
excl
A
)
:
Lemma
excl_validI
{
M
}
(
x
:
excl
A
)
:
(
✓
x
)%
I
≡
(
if
x
is
ExclBot
then
False
else
True
:
uPred
M
)%
I
.
(
✓
x
)%
I
≡
(
if
x
is
ExclBot
then
False
else
True
:
uPred
M
)%
I
.
Proof
.
by
destruct
x
.
Qed
.
Proof
.
uPred
.
unseal
.
by
destruct
x
.
Qed
.
(** ** Local updates *)
(** ** Local updates *)
Global
Instance
excl_local_update
b
:
Global
Instance
excl_local_update
b
:
...
...
algebra/fin_maps.v
View file @
2d9c5f33
...
@@ -171,9 +171,9 @@ Proof. split; [apply _|]. intros m ? i. by apply: cmra_discrete_valid. Qed.
...
@@ -171,9 +171,9 @@ Proof. split; [apply _|]. intros m ? i. by apply: cmra_discrete_valid. Qed.
(** Internalized properties *)
(** Internalized properties *)
Lemma
map_equivI
{
M
}
m1
m2
:
(
m1
≡
m2
)%
I
≡
(
∀
i
,
m1
!!
i
≡
m2
!!
i
:
uPred
M
)%
I
.
Lemma
map_equivI
{
M
}
m1
m2
:
(
m1
≡
m2
)%
I
≡
(
∀
i
,
m1
!!
i
≡
m2
!!
i
:
uPred
M
)%
I
.
Proof
.
done
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
Lemma
map_validI
{
M
}
m
:
(
✓
m
)%
I
≡
(
∀
i
,
✓
(
m
!!
i
)
:
uPred
M
)%
I
.
Lemma
map_validI
{
M
}
m
:
(
✓
m
)%
I
≡
(
∀
i
,
✓
(
m
!!
i
)
:
uPred
M
)%
I
.
Proof
.
done
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
End
cmra
.
End
cmra
.
Arguments
mapRA
_
{
_
_
}
_
.
Arguments
mapRA
_
{
_
_
}
_
.
...
...
algebra/iprod.v
View file @
2d9c5f33
...
@@ -171,9 +171,9 @@ Section iprod_cmra.
...
@@ -171,9 +171,9 @@ Section iprod_cmra.
(** Internalized properties *)
(** Internalized properties *)
Lemma
iprod_equivI
{
M
}
g1
g2
:
(
g1
≡
g2
)%
I
≡
(
∀
i
,
g1
i
≡
g2
i
:
uPred
M
)%
I
.
Lemma
iprod_equivI
{
M
}
g1
g2
:
(
g1
≡
g2
)%
I
≡
(
∀
i
,
g1
i
≡
g2
i
:
uPred
M
)%
I
.
Proof
.
done
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
Lemma
iprod_validI
{
M
}
g
:
(
✓
g
)%
I
≡
(
∀
i
,
✓
g
i
:
uPred
M
)%
I
.
Lemma
iprod_validI
{
M
}
g
:
(
✓
g
)%
I
≡
(
∀
i
,
✓
g
i
:
uPred
M
)%
I
.
Proof
.
done
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
(** Properties of iprod_insert. *)
(** Properties of iprod_insert. *)
Context
`
{
∀
x
x'
:
A
,
Decision
(
x
=
x'
)}.
Context
`
{
∀
x
x'
:
A
,
Decision
(
x
=
x'
)}.
...
...
algebra/option.v
View file @
2d9c5f33
...
@@ -140,10 +140,12 @@ Lemma option_equivI {M} (x y : option A) :
...
@@ -140,10 +140,12 @@ Lemma option_equivI {M} (x y : option A) :
(
x
≡
y
)%
I
≡
(
match
x
,
y
with
(
x
≡
y
)%
I
≡
(
match
x
,
y
with
|
Some
a
,
Some
b
=>
a
≡
b
|
None
,
None
=>
True
|
_
,
_
=>
False
|
Some
a
,
Some
b
=>
a
≡
b
|
None
,
None
=>
True
|
_
,
_
=>
False
end
:
uPred
M
)%
I
.
end
:
uPred
M
)%
I
.
Proof
.
do
2
split
.
by
destruct
1
.
by
destruct
x
,
y
;
try
constructor
.
Qed
.
Proof
.
uPred
.
unseal
.
do
2
split
.
by
destruct
1
.
by
destruct
x
,
y
;
try
constructor
.
Qed
.
Lemma
option_validI
{
M
}
(
x
:
option
A
)
:
Lemma
option_validI
{
M
}
(
x
:
option
A
)
:
(
✓
x
)%
I
≡
(
match
x
with
Some
a
=>
✓
a
|
None
=>
True
end
:
uPred
M
)%
I
.
(
✓
x
)%
I
≡
(
match
x
with
Some
a
=>
✓
a
|
None
=>
True
end
:
uPred
M
)%
I
.
Proof
.
by
destruct
x
.
Qed
.
Proof
.
uPred
.
unseal
.
by
destruct
x
.
Qed
.
(** Updates *)
(** Updates *)
Lemma
option_updateP
(
P
:
A
→
Prop
)
(
Q
:
option
A
→
Prop
)
x
:
Lemma
option_updateP
(
P
:
A
→
Prop
)
(
Q
:
option
A
→
Prop
)
x
:
...
...
algebra/upred.v
View file @
2d9c5f33
This diff is collapsed.
Click to expand it.
program_logic/adequacy.v
View file @
2d9c5f33
...
@@ -62,8 +62,8 @@ Proof.
...
@@ -62,8 +62,8 @@ Proof.
intros
Hht
????
;
apply
(
nsteps_wptp
[
Φ
]
k
n
([
e1
],
σ
1
)
(
t2
,
σ
2
)
[
r1
])
;
intros
Hht
????
;
apply
(
nsteps_wptp
[
Φ
]
k
n
([
e1
],
σ
1
)
(
t2
,
σ
2
)
[
r1
])
;
rewrite
/
big_op
?right_id
;
auto
.
rewrite
/
big_op
?right_id
;
auto
.
constructor
;
last
constructor
.
constructor
;
last
constructor
.
apply
Hht
with
(
k
+
n
)
r1
;
eauto
using
cmra_includ
ed
_
un
i
t
.
move
:
Hht
;
rewrite
/
ht
;
uPr
ed
.
un
seal
=>
Hh
t
.
e
apply
uPred
.
const_intro
;
eauto
.
apply
Hht
with
(
k
+
n
)
r1
;
by
eauto
using
cmra_included_unit
.
Qed
.
Qed
.
Lemma
ht_adequacy_own
Φ
e1
t2
σ
1
m
σ
2
:
Lemma
ht_adequacy_own
Φ
e1
t2
σ
1
m
σ
2
:
✓
m
→
✓
m
→
...
@@ -74,9 +74,9 @@ Proof.
...
@@ -74,9 +74,9 @@ Proof.
intros
Hv
?
[
k
?]%
rtc_nsteps
.
intros
Hv
?
[
k
?]%
rtc_nsteps
.
eapply
ht_adequacy_steps
with
(
r1
:
=
(
Res
∅
(
Excl
σ
1
)
(
Some
m
)))
;
eauto
;
[|].
eapply
ht_adequacy_steps
with
(
r1
:
=
(
Res
∅
(
Excl
σ
1
)
(
Some
m
)))
;
eauto
;
[|].
{
by
rewrite
Nat
.
add_comm
;
apply
wsat_init
,
cmra_valid_validN
.
}
{
by
rewrite
Nat
.
add_comm
;
apply
wsat_init
,
cmra_valid_validN
.
}
exists
(
Res
∅
(
Excl
σ
1
)
∅
),
(
Res
∅
∅
(
Some
m
))
;
split_and
?.
uPred
.
unseal
;
exists
(
Res
∅
(
Excl
σ
1
)
∅
),
(
Res
∅
∅
(
Some
m
))
;
split_and
?.
-
by
rewrite
Res_op
?left_id
?right_id
.
-
by
rewrite
Res_op
?left_id
?right_id
.
-
by
rewrite
/
uPred_holds
/=.
-
rewrite
/
ownP
;
uPred
.
unseal
;
rewrite
/
uPred_holds
/
/=.
-
by
apply
ownG_spec
.
-
by
apply
ownG_spec
.
Qed
.
Qed
.
Theorem
ht_adequacy_result
E
φ
e
v
t2
σ
1
m
σ
2
:
Theorem
ht_adequacy_result
E
φ
e
v
t2
σ
1
m
σ
2
:
...
@@ -90,8 +90,8 @@ Proof.
...
@@ -90,8 +90,8 @@ Proof.
as
(
rs2
&
Qs
&
Hwptp
&?)
;
auto
.
as
(
rs2
&
Qs
&
Hwptp
&?)
;
auto
.
{
by
rewrite
-(
ht_mask_weaken
E
⊤
).
}
{
by
rewrite
-(
ht_mask_weaken
E
⊤
).
}
inversion
Hwptp
as
[|??
r
??
rs
Hwp
_
]
;
clear
Hwptp
;
subst
.
inversion
Hwptp
as
[|??
r
??
rs
Hwp
_
]
;
clear
Hwptp
;
subst
.
apply
wp_value_inv
in
Hwp
;
destruct
(
Hwp
(
big_op
rs
)
2
∅
σ
2
)
as
[
r'
[]]
;
auto
.
move
:
Hwp
.
uPred
.
unseal
=>
/
wp_value_inv
Hwp
.
by
rewrite
right_id_L
.
destruct
(
Hwp
(
big_op
rs
)
2
∅
σ
2
)
as
[
r'
[]]
;
rewrite
?
right_id_L
;
auto
.
Qed
.
Qed
.
Lemma
ht_adequacy_reducible
E
Φ
e1
e2
t2
σ
1
m
σ
2
:
Lemma
ht_adequacy_reducible
E
Φ
e1
e2
t2
σ
1
m
σ
2
:
✓
m
→
✓
m
→
...
...
program_logic/ghost_ownership.v
View file @
2d9c5f33
...
@@ -102,7 +102,7 @@ Proof.
...
@@ -102,7 +102,7 @@ Proof.
first
(
eapply
map_updateP_alloc_strong'
,
cmra_transport_valid
,
Ha
)
;
first
(
eapply
map_updateP_alloc_strong'
,
cmra_transport_valid
,
Ha
)
;
naive_solver
.
naive_solver
.
-
apply
exist_elim
=>
m
;
apply
const_elim_l
=>-[
γ
[
Hfresh
->]].
-
apply
exist_elim
=>
m
;
apply
const_elim_l
=>-[
γ
[
Hfresh
->]].
by
rewrite
-(
exist_intro
γ
)
const_equiv
.
by
rewrite
-(
exist_intro
γ
)
const_equiv
//
left_id
.
Qed
.
Qed
.
Lemma
own_alloc
a
E
:
✓
a
→
True
⊑
(|={
E
}=>
∃
γ
,
own
γ
a
).
Lemma
own_alloc
a
E
:
✓
a
→
True
⊑
(|={
E
}=>
∃
γ
,
own
γ
a
).
Proof
.
Proof
.
...
...
program_logic/lifting.v
View file @
2d9c5f33
...
@@ -16,7 +16,7 @@ Implicit Types P Q : iProp Λ Σ.
...
@@ -16,7 +16,7 @@ Implicit Types P Q : iProp Λ Σ.
Implicit
Types
Φ
:
val
Λ
→
iProp
Λ
Σ
.
Implicit
Types
Φ
:
val
Λ
→
iProp
Λ
Σ
.
Transparent
uPred_holds
.
Transparent
uPred_holds
.
Notation
wp_fork
ef
:
=
(
default
True
ef
(
flip
(
wp
⊤
)
(
λ
_
,
True
)))%
I
.
Notation
wp_fork
ef
:
=
(
default
True
ef
(
flip
(
wp
⊤
)
(
λ
_
,
■
True
)))%
I
.
Lemma
wp_lift_step
E1
E2
Lemma
wp_lift_step
E1
E2
(
φ
:
expr
Λ
→
state
Λ
→
option
(
expr
Λ
)
→
Prop
)
Φ
e1
σ
1
:
(
φ
:
expr
Λ
→
state
Λ
→
option
(
expr
Λ
)
→
Prop
)
Φ
e1
σ
1
:
...
@@ -27,7 +27,7 @@ Lemma wp_lift_step E1 E2
...
@@ -27,7 +27,7 @@ Lemma wp_lift_step E1 E2
(
■
φ
e2
σ
2
ef
∧
ownP
σ
2
)
-
★
|={
E1
,
E2
}=>
||
e2
@
E2
{{
Φ
}}
★
wp_fork
ef
)
(
■
φ
e2
σ
2
ef
∧
ownP
σ
2
)
-
★
|={
E1
,
E2
}=>
||
e2
@
E2
{{
Φ
}}
★
wp_fork
ef
)
⊑
||
e1
@
E2
{{
Φ
}}.
⊑
||
e1
@
E2
{{
Φ
}}.
Proof
.
Proof
.
intros
?
He
Hsafe
Hstep
;
split
=>
n
r
?
Hvs
;
constructor
;
auto
.
intros
?
He
Hsafe
Hstep
.
uPred
.
unseal
;
split
=>
n
r
?
Hvs
;
constructor
;
auto
.
intros
rf
k
Ef
σ
1
'
???
;
destruct
(
Hvs
rf
(
S
k
)
Ef
σ
1
'
)
intros
rf
k
Ef
σ
1
'
???
;
destruct
(
Hvs
rf
(
S
k
)
Ef
σ
1
'
)
as
(
r'
&(
r1
&
r2
&?&?&
Hwp
)&
Hws
)
;
auto
;
clear
Hvs
;
cofe_subst
r'
.
as
(
r'
&(
r1
&
r2
&?&?&
Hwp
)&
Hws
)
;
auto
;
clear
Hvs
;
cofe_subst
r'
.
destruct
(
wsat_update_pst
k
(
E1
∪
Ef
)
σ
1
σ
1
'
r1
(
r2
⋅
rf
))
as
[->
Hws'
].
destruct
(
wsat_update_pst
k
(
E1
∪
Ef
)
σ
1
σ
1
'
r1
(
r2
⋅
rf
))
as
[->
Hws'
].
...
@@ -38,7 +38,7 @@ Proof.
...
@@ -38,7 +38,7 @@ Proof.
as
(
r'
&(
r1'
&
r2'
&?&?&?)&?)
;
auto
;
cofe_subst
r'
.
as
(
r'
&(
r1'
&
r2'
&?&?&?)&?)
;
auto
;
cofe_subst
r'
.
{
split
.
by
eapply
Hstep
.
apply
ownP_spec
;
auto
.
}
{
split
.
by
eapply
Hstep
.
apply
ownP_spec
;
auto
.
}
{
rewrite
(
comm
_
r2
)
-
assoc
;
eauto
using
wsat_le
.
}
{
rewrite
(
comm
_
r2
)
-
assoc
;
eauto
using
wsat_le
.
}
by
exists
r1'
,
r2'
;
split_and
?
;
[|
|
by
intros
?
->
]
.
exists
r1'
,
r2'
;
split_and
?
;
try
done
.
by
uPred
.
unseal
;
intros
?
->.
Qed
.
Qed
.
Lemma
wp_lift_pure_step
E
(
φ
:
expr
Λ
→
option
(
expr
Λ
)
→
Prop
)
Φ
e1
:
Lemma
wp_lift_pure_step
E
(
φ
:
expr
Λ
→
option
(
expr
Λ
)
→
Prop
)
Φ
e1
:
...
@@ -47,11 +47,13 @@ Lemma wp_lift_pure_step E (φ : expr Λ → option (expr Λ) → Prop) Φ e1 :
...
@@ -47,11 +47,13 @@ Lemma wp_lift_pure_step E (φ : expr Λ → option (expr Λ) → Prop) Φ e1 :
(
∀
σ
1 e2
σ
2
ef
,
prim_step
e1
σ
1 e2
σ
2
ef
→
σ
1
=
σ
2
∧
φ
e2
ef
)
→
(
∀
σ
1 e2
σ
2
ef
,
prim_step
e1
σ
1 e2
σ
2
ef
→
σ
1
=
σ
2
∧
φ
e2
ef
)
→
(
▷
∀
e2
ef
,
■
φ
e2
ef
→
||
e2
@
E
{{
Φ
}}
★
wp_fork
ef
)
⊑
||
e1
@
E
{{
Φ
}}.
(
▷
∀
e2
ef
,
■
φ
e2
ef
→
||
e2
@
E
{{
Φ
}}
★
wp_fork
ef
)
⊑
||
e1
@
E
{{
Φ
}}.
Proof
.
Proof
.
intros
He
Hsafe
Hstep
;
split
=>
n
r
?
Hwp
;
constructor
;
auto
.
intros
He
Hsafe
Hstep
;
uPred
.
unseal
;
split
=>
n
r
?
Hwp
;
constructor
;
auto
.
intros
rf
k
Ef
σ
1
???
;
split
;
[
done
|].
destruct
n
as
[|
n
]
;
first
lia
.
intros
rf
k
Ef
σ
1
???
;
split
;
[
done
|].
destruct
n
as
[|
n
]
;
first
lia
.
intros
e2
σ
2
ef
?
;
destruct
(
Hstep
σ
1 e2
σ
2
ef
)
;
auto
;
subst
.
intros
e2
σ
2
ef
?
;
destruct
(
Hstep
σ
1 e2
σ
2
ef
)
;
auto
;
subst
.
destruct
(
Hwp
e2
ef
k
r
)
as
(
r1
&
r2
&
Hr
&?&?)
;
auto
.
destruct
(
Hwp
e2
ef
k
r
)
as
(
r1
&
r2
&
Hr
&?&?)
;
auto
.
exists
r1
,
r2
;
split_and
?
;
[
rewrite
-
Hr
|
|
by
intros
?
->]
;
eauto
using
wsat_le
.
exists
r1
,
r2
;
split_and
?
;
try
done
.
-
rewrite
-
Hr
;
eauto
using
wsat_le
.
-
uPred
.
unseal
;
by
intros
?
->.
Qed
.
Qed
.
(** Derived lifting lemmas. *)
(** Derived lifting lemmas. *)
...
...
program_logic/ownership.v
View file @
2d9c5f33
...
@@ -75,7 +75,8 @@ Lemma ownI_spec n r i P :
...
@@ -75,7 +75,8 @@ Lemma ownI_spec n r i P :
✓
{
n
}
r
→
✓
{
n
}
r
→
(
ownI
i
P
)
n
r
↔
wld
r
!!
i
≡
{
n
}
≡
Some
(
to_agree
(
Next
(
iProp_unfold
P
))).
(
ownI
i
P
)
n
r
↔
wld
r
!!
i
≡
{
n
}
≡
Some
(
to_agree
(
Next
(
iProp_unfold
P
))).
Proof
.
Proof
.
intros
[??]
;
rewrite
/
uPred_holds
/=
res_includedN
/=
singleton_includedN
;
split
.
intros
(?&?&?).
rewrite
/
ownI
;
uPred
.
unseal
.
rewrite
/
uPred_holds
/=
res_includedN
/=
singleton_includedN
;
split
.
-
intros
[(
P'
&
Hi
&
HP
)
_
]
;
rewrite
Hi
.
-
intros
[(
P'
&
Hi
&
HP
)
_
]
;
rewrite
Hi
.
by
apply
Some_dist
,
symmetry
,
agree_valid_includedN
,
by
apply
Some_dist
,
symmetry
,
agree_valid_includedN
,
(
cmra_included_includedN
_
P'
),
HP
;
apply
map_lookup_validN
with
(
wld
r
)
i
.
(
cmra_included_includedN
_
P'
),
HP
;
apply
map_lookup_validN
with
(
wld
r
)
i
.
...
@@ -83,11 +84,13 @@ Proof.
...
@@ -83,11 +84,13 @@ Proof.
Qed
.
Qed
.
Lemma
ownP_spec
n
r
σ
:
✓
{
n
}
r
→
(
ownP
σ
)
n
r
↔
pst
r
≡
Excl
σ
.
Lemma
ownP_spec
n
r
σ
:
✓
{
n
}
r
→
(
ownP
σ
)
n
r
↔
pst
r
≡
Excl
σ
.
Proof
.
Proof
.
intros
(?&?&?)
;
rewrite
/
uPred_holds
/=
res_includedN
/=
Excl_includedN
//.
intros
(?&?&?).
rewrite
/
ownP
;
uPred
.
unseal
.
rewrite
/
uPred_holds
/=
res_includedN
/=
Excl_includedN
//.
rewrite
(
timeless_iff
n
).
naive_solver
(
apply
cmra_empty_leastN
).
rewrite
(
timeless_iff
n
).
naive_solver
(
apply
cmra_empty_leastN
).
Qed
.
Qed
.
Lemma
ownG_spec
n
r
m
:
(
ownG
m
)
n
r
↔
Some
m
≼
{
n
}
gst
r
.
Lemma
ownG_spec
n
r
m
:
(
ownG
m
)
n
r
↔
Some
m
≼
{
n
}
gst
r
.
Proof
.
Proof
.
rewrite
/
ownG
;
uPred
.
unseal
.
rewrite
/
uPred_holds
/=
res_includedN
;
naive_solver
(
apply
cmra_empty_leastN
).
rewrite
/
uPred_holds
/=
res_includedN
;
naive_solver
(
apply
cmra_empty_leastN
).
Qed
.
Qed
.
End
ownership
.
End
ownership
.
program_logic/pviewshifts.v
View file @
2d9c5f33
...
@@ -63,7 +63,7 @@ Qed.
...
@@ -63,7 +63,7 @@ Qed.
Lemma
pvs_timeless
E
P
:
TimelessP
P
→
(
▷
P
)
⊑
(|={
E
}=>
P
).
Lemma
pvs_timeless
E
P
:
TimelessP
P
→
(
▷
P
)
⊑
(|={
E
}=>
P
).
Proof
.
Proof
.
rewrite
uPred
.
timelessP_spec
=>
HP
.
rewrite
uPred
.
timelessP_spec
=>
HP
.
split
=>-[|
n
]
r
?
HP'
rf
k
Ef
σ
???
;
first
lia
.
uPred
.
unseal
;
split
=>-[|
n
]
r
?
HP'
rf
k
Ef
σ
???
;
first
lia
.
exists
r
;
split
;
last
done
.
exists
r
;
split
;
last
done
.
apply
HP
,
uPred_weaken
with
n
r
;
eauto
using
cmra_validN_le
.
apply
HP
,
uPred_weaken
with
n
r
;
eauto
using
cmra_validN_le
.
Qed
.
Qed
.
...
@@ -82,7 +82,7 @@ Proof.
...
@@ -82,7 +82,7 @@ Proof.
Qed
.
Qed
.
Lemma
pvs_frame_r
E1
E2
P
Q
:
((|={
E1
,
E2
}=>
P
)
★
Q
)
⊑
(|={
E1
,
E2
}=>
P
★
Q
).
Lemma
pvs_frame_r
E1
E2
P
Q
:
((|={
E1
,
E2
}=>
P
)
★
Q
)
⊑
(|={
E1
,
E2
}=>
P
★
Q
).
Proof
.
Proof
.
split
;
intros
n
r
?
(
r1
&
r2
&
Hr
&
HP
&?)
rf
k
Ef
σ
???.
uPred
.
unseal
;
split
;
intros
n
r
?
(
r1
&
r2
&
Hr
&
HP
&?)
rf
k
Ef
σ
???.
destruct
(
HP
(
r2
⋅
rf
)
k
Ef
σ
)
as
(
r'
&?&?)
;
eauto
.
destruct
(
HP
(
r2
⋅
rf
)
k
Ef
σ
)
as
(
r'
&?&?)
;
eauto
.
{
by
rewrite
assoc
-(
dist_le
_
_
_
_
Hr
)
;
last
lia
.
}
{
by
rewrite
assoc
-(
dist_le
_
_
_
_
Hr
)
;
last
lia
.
}
exists
(
r'
⋅
r2
)
;
split
;
last
by
rewrite
-
assoc
.
exists
(
r'
⋅
r2
)
;
split
;
last
by
rewrite
-
assoc
.
...
@@ -90,7 +90,7 @@ Proof.
...
@@ -90,7 +90,7 @@ Proof.
Qed
.
Qed
.
Lemma
pvs_openI
i
P
:
ownI
i
P
⊑
(|={{[
i
]},
∅
}=>
▷
P
).
Lemma
pvs_openI
i
P
:
ownI
i
P
⊑
(|={{[
i
]},
∅
}=>
▷
P
).
Proof
.
Proof
.
split
=>
-[|
n
]
r
?
Hinv
rf
[|
k
]
Ef
σ
???
;
try
lia
.
uPred
.
unseal
;
split
=>
-[|
n
]
r
?
Hinv
rf
[|
k
]
Ef
σ
???
;
try
lia
.
apply
ownI_spec
in
Hinv
;
last
auto
.
apply
ownI_spec
in
Hinv
;
last
auto
.
destruct
(
wsat_open
k
Ef
σ
(
r
⋅
rf
)
i
P
)
as
(
rP
&?&?)
;
auto
.
destruct
(
wsat_open
k
Ef
σ
(
r
⋅
rf
)
i
P
)
as
(
rP
&?&?)
;
auto
.
{
rewrite
lookup_wld_op_l
?Hinv
;
eauto
;
apply
dist_le
with
(
S
n
)
;
eauto
.
}
{
rewrite
lookup_wld_op_l
?Hinv
;
eauto
;
apply
dist_le
with
(
S
n
)
;
eauto
.
}
...
@@ -99,7 +99,7 @@ Proof.
...
@@ -99,7 +99,7 @@ Proof.
Qed
.
Qed
.
Lemma
pvs_closeI
i
P
:
(
ownI
i
P
∧
▷
P
)
⊑
(|={
∅
,{[
i
]}}=>
True
).
Lemma
pvs_closeI
i
P
:
(
ownI
i
P
∧
▷
P
)
⊑
(|={
∅
,{[
i
]}}=>
True
).
Proof
.
Proof
.
split
=>
-[|
n
]
r
?
[?
HP
]
rf
[|
k
]
Ef
σ
?
HE
?
;
try
lia
.
uPred
.
unseal
;
split
=>
-[|
n
]
r
?
[?
HP
]
rf
[|
k
]
Ef
σ
?
HE
?
;
try
lia
.
exists
∅
;
split
;
[
done
|].
exists
∅
;
split
;
[
done
|].
rewrite
left_id
;
apply
wsat_close
with
P
r
.
rewrite
left_id
;
apply
wsat_close
with
P
r
.
-
apply
ownI_spec
,
uPred_weaken
with
(
S
n
)
r
;
auto
.
-
apply
ownI_spec
,
uPred_weaken
with
(
S
n
)
r
;
auto
.
...
@@ -111,7 +111,7 @@ Lemma pvs_ownG_updateP E m (P : iGst Λ Σ → Prop) :
...
@@ -111,7 +111,7 @@ Lemma pvs_ownG_updateP E m (P : iGst Λ Σ → Prop) :
m
~~>
:
P
→
ownG
m
⊑
(|={
E
}=>
∃
m'
,
■
P
m'
∧
ownG
m'
).
m
~~>
:
P
→
ownG
m
⊑
(|={
E
}=>
∃
m'
,
■
P
m'
∧
ownG
m'
).
Proof
.
Proof
.
intros
Hup
%
option_updateP'
.
intros
Hup
%
option_updateP'
.
split
=>
-[|
n
]
r
?
/
ownG_spec
Hinv
rf
[|
k
]
Ef
σ
???
;
try
lia
.
uPred
.
unseal
;
split
=>
-[|
n
]
r
?
/
ownG_spec
Hinv
rf
[|
k
]
Ef
σ
???
;
try
lia
.
destruct
(
wsat_update_gst
k
(
E
∪
Ef
)
σ
r
rf
(
Some
m
)
P
)
as
(
m'
&?&?)
;
eauto
.
destruct
(
wsat_update_gst
k
(
E
∪
Ef
)
σ
r
rf
(
Some
m
)
P
)
as
(
m'
&?&?)
;
eauto
.
{
apply
cmra_includedN_le
with
(
S
n
)
;
auto
.
}
{
apply
cmra_includedN_le
with
(
S
n
)
;
auto
.
}
by
exists
(
update_gst
m'
r
)
;
split
;
[
exists
m'
;
split
;
[|
apply
ownG_spec
]|].
by
exists
(
update_gst
m'
r
)
;
split
;
[
exists
m'
;
split
;
[|
apply
ownG_spec
]|].
...
@@ -120,20 +120,21 @@ Lemma pvs_ownG_updateP_empty `{Empty (iGst Λ Σ), !CMRAIdentity (iGst Λ Σ)}
...
@@ -120,20 +120,21 @@ Lemma pvs_ownG_updateP_empty `{Empty (iGst Λ Σ), !CMRAIdentity (iGst Λ Σ)}
E
(
P
:
iGst
Λ
Σ
→
Prop
)
:
E
(
P
:
iGst
Λ
Σ
→
Prop
)
:
∅
~~>
:
P
→
True
⊑
(|={
E
}=>
∃
m'
,
■
P
m'
∧
ownG
m'
).
∅
~~>
:
P
→
True
⊑
(|={
E
}=>
∃
m'
,
■
P
m'
∧
ownG
m'
).
Proof
.
Proof
.
intros
Hup
;
split
=>
-[|
n
]
r
?
_
rf
[|
k
]
Ef
σ
???
;
try
lia
.
intros
Hup
;
uPred
.
unseal
;
split
=>
-[|
n
]
r
?
_
rf
[|
k
]
Ef
σ
???
;
try
lia
.
destruct
(
wsat_update_gst
k
(
E
∪
Ef
)
σ
r
rf
∅
P
)
as
(
m'
&?&?)
;
eauto
.
destruct
(
wsat_update_gst
k
(
E
∪
Ef
)
σ
r
rf
∅
P
)
as
(
m'
&?&?)
;
eauto
.
{
apply
cmra_empty_leastN
.
}
{
apply
cmra_empty_leastN
.
}
{
apply
cmra_updateP_compose_l
with
(
Some
∅
),
option_updateP
with
P
;
{
apply
cmra_updateP_compose_l
with
(
Some
∅
),
option_updateP
with
P
;
auto
using
option_update_None
.
}
auto
using
option_update_None
.
}
by
exists
(
update_gst
m'
r
)
;
split
;
[
exists
m'
;
split
;
[|
apply
ownG_spec
]|].
exists
(
update_gst
m'
r
)
;
by
split
;
[
exists
m'
;
split
;
[|
apply
ownG_spec
]|].
Qed
.
Qed
.
Lemma
pvs_allocI
E
P
:
¬
set_finite
E
→
▷
P
⊑
(|={
E
}=>
∃
i
,
■
(
i
∈
E
)
∧
ownI
i
P
).
Lemma
pvs_allocI
E
P
:
¬
set_finite
E
→
▷
P
⊑
(|={
E
}=>
∃
i
,
■
(
i
∈
E
)
∧
ownI
i
P
).
Proof
.
Proof
.
intros
?
;
split
=>
-[|
n
]
r
?
HP
rf
[|
k
]
Ef
σ
???
;
try
lia
.
intros
?
;
rewrite
/
ownI
;
uPred
.
unseal
.
split
=>
-[|
n
]
r
?
HP
rf
[|
k
]
Ef
σ
???
;
try
lia
.
destruct
(
wsat_alloc
k
E
Ef
σ
rf
P
r
)
as
(
i
&?&?&?)
;
auto
.
destruct
(
wsat_alloc
k
E
Ef
σ
rf
P
r
)
as
(
i
&?&?&?)
;
auto
.
{
apply
uPred_weaken
with
n
r
;
eauto
.
}
{
apply
uPred_weaken
with
n
r
;
eauto
.
}
exists
(
Res
{[
i
:
=
to_agree
(
Next
(
iProp_unfold
P
))
]}
∅
∅
).
exists
(
Res
{[
i
:
=
to_agree
(
Next
(
iProp_unfold
P
))
]}
∅
∅
).
by
split
;
[
by
exists
i
;
split
;
rewrite
/
uPred_holds
/=
|]
.
split
;
[
|
done
].
by
exists
i
;
split
;
rewrite
/
uPred_holds
/=.
Qed
.
Qed
.
(** * Derived rules *)
(** * Derived rules *)
...
...
program_logic/resources.v
View file @
2d9c5f33
...
@@ -162,9 +162,11 @@ Proof. by intros ? ? [???]; constructor; apply: timeless. Qed.
...
@@ -162,9 +162,11 @@ Proof. by intros ? ? [???]; constructor; apply: timeless. Qed.
(** Internalized properties *)
(** Internalized properties *)
Lemma
res_equivI
{
M
}
r1
r2
:
Lemma
res_equivI
{
M
}
r1
r2
:
(
r1
≡
r2
)%
I
≡
(
wld
r1
≡
wld
r2
∧
pst
r1
≡
pst
r2
∧
gst
r1
≡
gst
r2
:
uPred
M
)%
I
.
(
r1
≡
r2
)%
I
≡
(
wld
r1
≡
wld
r2
∧
pst
r1
≡
pst
r2
∧
gst
r1
≡
gst
r2
:
uPred
M
)%
I
.
Proof
.
do
2
split
.
by
destruct
1
.
by
intros
(?&?&?)
;
constructor
.
Qed
.
Proof
.
uPred
.
unseal
.
do
2
split
.
by
destruct
1
.
by
intros
(?&?&?)
;
by
constructor
.
Qed
.
Lemma
res_validI
{
M
}
r
:
(
✓
r
)%
I
≡
(
✓
wld
r
∧
✓
pst
r
∧
✓
gst
r
:
uPred
M
)%
I
.
Lemma
res_validI
{
M
}
r
:
(
✓
r
)%
I
≡
(
✓
wld
r
∧
✓
pst
r
∧
✓
gst
r
:
uPred
M
)%
I
.
Proof
.
done
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
End
res
.
End
res
.
Arguments
resC
:
clear
implicits
.
Arguments
resC
:
clear
implicits
.
...
...
program_logic/sts.v
View file @
2d9c5f33
...
@@ -179,7 +179,7 @@ Section sts.
...
@@ -179,7 +179,7 @@ Section sts.
(
sts_own
γ
s'
T'
-
★
Ψ
x
)))
→
(
sts_own
γ
s'
T'
-
★
Ψ
x
)))
→
P
⊑
fsa
E
Ψ
.
P
⊑
fsa
E
Ψ
.
Proof
.
Proof
.
rewrite
sts_own_eq
.
intros
.
eapply
sts_fsaS
;
try
done
;
[].
(* FIXME: slow *)
rewrite
sts_own_eq
.
intros
.
eapply
sts_fsaS
;
try
done
;
[].
by
rewrite
sts_ownS_eq
sts_own_eq
.
by
rewrite
sts_ownS_eq
sts_own_eq
.
Qed
.
Qed
.
End
sts
.
End
sts
.
program_logic/weakestpre.v
View file @
2d9c5f33
...
@@ -158,11 +158,12 @@ Proof.
...
@@ -158,11 +158,12 @@ Proof.
Qed
.
Qed
.
Lemma
wp_frame_r
E
e
Φ
R
:
(||
e
@
E
{{
Φ
}}
★
R
)
⊑
||
e
@
E
{{
λ
v
,
Φ
v
★
R
}}.
Lemma
wp_frame_r
E
e
Φ
R
:
(||
e
@
E
{{
Φ
}}
★
R
)
⊑
||
e
@
E
{{
λ
v
,
Φ
v
★
R
}}.
Proof
.
Proof
.
split
;
intros
n
r'
Hvalid
(
r
&
rR
&
Hr
&
Hwp
&?)
;
revert
Hvalid
.
uPred
.
unseal
;
split
;
intros
n
r'
Hvalid
(
r
&
rR
&
Hr
&
Hwp
&?)
;
revert
Hvalid
.
rewrite
Hr
;
clear
Hr
;
revert
e
r
Hwp
.
rewrite
Hr
;
clear
Hr
;
revert
e
r
Hwp
.
induction
n
as
[
n
IH
]
using
lt_wf_ind
;
intros
e
r1
.
induction
n
as
[
n
IH
]
using
lt_wf_ind
;
intros
e
r1
.
destruct
1
as
[|
n
r
e
?
Hgo
]=>?.
destruct
1
as
[|
n
r
e
?
Hgo
]=>?.
{
constructor
;
apply
pvs_frame_r
;
auto
.
exists
r
,
rR
;
eauto
.
}
{
constructor
.
rewrite
-
uPred_sep_eq
;
apply
pvs_frame_r
;
auto
.
uPred
.
unseal
;
exists
r
,
rR
;
eauto
.
}
constructor
;
[
done
|]=>
rf
k
Ef
σ
1
???.
constructor
;
[
done
|]=>
rf
k
Ef
σ
1
???.
destruct
(
Hgo
(
rR
⋅
rf
)
k
Ef
σ
1
)
as
[
Hsafe
Hstep
]
;
auto
.
destruct
(
Hgo
(
rR
⋅
rf
)
k
Ef
σ
1
)
as
[
Hsafe
Hstep
]
;
auto
.
{
by
rewrite
assoc
.
}
{
by
rewrite
assoc
.
}
...
@@ -176,7 +177,7 @@ Qed.
...
@@ -176,7 +177,7 @@ Qed.
Lemma
wp_frame_later_r
E
e
Φ
R
:
Lemma
wp_frame_later_r
E
e
Φ
R
:
to_val
e
=
None
→
(||
e
@
E
{{
Φ
}}
★
▷
R
)
⊑
||
e
@
E
{{
λ
v
,
Φ
v
★
R
}}.
to_val
e
=
None
→
(||
e
@
E
{{
Φ
}}
★
▷
R
)
⊑
||
e
@
E
{{
λ
v
,
Φ
v
★
R
}}.
Proof
.
Proof
.
intros
He
;
split
;
intros
n
r'
Hvalid
(
r
&
rR
&
Hr
&
Hwp
&?).
intros
He
;
uPred
.
unseal
;
split
;
intros
n
r'
Hvalid
(
r
&
rR
&
Hr
&
Hwp
&?).
revert
Hvalid
;
rewrite
Hr
;
clear
Hr
.
revert
Hvalid
;
rewrite
Hr
;
clear
Hr
.
destruct
Hwp
as
[|
n
r
e
?
Hgo
]
;
[
by
rewrite
to_of_val
in
He
|].
destruct
Hwp
as
[|
n
r
e
?
Hgo
]
;
[
by
rewrite
to_of_val
in
He
|].
constructor
;
[
done
|]=>
rf
k
Ef
σ
1
???
;
destruct
n
as
[|
n
]
;
first
omega
.
constructor
;
[
done
|]=>
rf
k
Ef
σ
1
???
;
destruct
n
as
[|
n
]
;
first
omega
.
...
@@ -185,7 +186,8 @@ Proof.
...
@@ -185,7 +186,8 @@ Proof.
destruct
(
Hstep
e2
σ
2
ef
)
as
(
r2
&
r2'
&?&?&?)
;
auto
.
destruct
(
Hstep
e2
σ
2
ef
)
as
(
r2
&
r2'
&?&?&?)
;
auto
.
exists
(
r2
⋅
rR
),
r2'
;
split_and
?
;
auto
.
exists
(
r2
⋅
rR
),
r2'
;
split_and
?
;
auto
.
-
by
rewrite
-(
assoc
_
r2
)
(
comm
_
rR
)
!
assoc
-(
assoc
_
_
rR
).
-
by
rewrite
-(
assoc
_
r2
)
(
comm
_
rR
)
!
assoc
-(
assoc
_
_
rR
).
-
apply
wp_frame_r
;
[
auto
|
exists
r2
,
rR
;
split_and
?
;
auto
].
-
rewrite
-
uPred_sep_eq
.
apply
wp_frame_r
;
[
auto
|
uPred
.
unseal
;
exists
r2
,
rR
;
split_and
?
;
auto
].
eapply
uPred_weaken
with
n
rR
;
eauto
.
eapply
uPred_weaken
with
n
rR
;
eauto
.
Qed
.
Qed
.
Lemma
wp_bind
`
{
LanguageCtx
Λ
K
}
E
e
Φ
:
Lemma
wp_bind
`
{
LanguageCtx
Λ
K
}
E
e
Φ
:
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment