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
Rodolphe Lepigre
Iris
Commits
b4567fbd
Commit
b4567fbd
authored
Oct 22, 2017
by
Robbert Krebbers
Browse files
Rename `always` → `persistently` (the persistent modality).
parent
0ad1d2bd
Changes
15
Expand all
Hide whitespace changes
Inline
Side-by-side
theories/base_logic/base_logic.v
View file @
b4567fbd
...
...
@@ -11,7 +11,7 @@ End uPred.
Hint
Resolve
pure_intro
.
Hint
Resolve
or_elim
or_intro_l'
or_intro_r'
:
I
.
Hint
Resolve
and_intro
and_elim_l'
and_elim_r'
:
I
.
Hint
Resolve
always
_mono
:
I
.
Hint
Resolve
persistently
_mono
:
I
.
Hint
Resolve
sep_elim_l'
sep_elim_r'
sep_mono
:
I
.
Hint
Immediate
True_intro
False_elim
:
I
.
Hint
Immediate
iff_refl
internal_eq_refl'
:
I
.
theories/base_logic/big_op.v
View file @
b4567fbd
...
...
@@ -117,11 +117,11 @@ Section list.
▷
^
n
([
∗
list
]
k
↦
x
∈
l
,
Φ
k
x
)
⊣
⊢
([
∗
list
]
k
↦
x
∈
l
,
▷
^
n
Φ
k
x
).
Proof
.
apply
(
big_opL_commute
_
).
Qed
.
Lemma
big_sepL_
always
Φ
l
:
Lemma
big_sepL_
persistently
Φ
l
:
(
□
[
∗
list
]
k
↦
x
∈
l
,
Φ
k
x
)
⊣
⊢
([
∗
list
]
k
↦
x
∈
l
,
□
Φ
k
x
).
Proof
.
apply
(
big_opL_commute
_
).
Qed
.
Lemma
big_sepL_
always
_if
p
Φ
l
:
Lemma
big_sepL_
persistently
_if
p
Φ
l
:
□
?p
([
∗
list
]
k
↦
x
∈
l
,
Φ
k
x
)
⊣
⊢
([
∗
list
]
k
↦
x
∈
l
,
□
?p
Φ
k
x
).
Proof
.
apply
(
big_opL_commute
_
).
Qed
.
...
...
@@ -134,7 +134,7 @@ Section list.
apply
impl_intro_l
,
pure_elim_l
=>
?
;
by
apply
big_sepL_lookup
.
}
revert
Φ
H
Φ
.
induction
l
as
[|
x
l
IH
]=>
Φ
H
Φ
.
{
rewrite
big_sepL_nil
;
auto
with
I
.
}
rewrite
big_sepL_cons
.
rewrite
-
always
_and_sep_l
;
apply
and_intro
.
rewrite
big_sepL_cons
.
rewrite
-
persistently
_and_sep_l
;
apply
and_intro
.
-
by
rewrite
(
forall_elim
0
)
(
forall_elim
x
)
pure_True
//
True_impl
.
-
rewrite
-
IH
.
apply
forall_intro
=>
k
;
by
rewrite
(
forall_elim
(
S
k
)).
Qed
.
...
...
@@ -143,10 +143,10 @@ Section list.
□
(
∀
k
x
,
⌜
l
!!
k
=
Some
x
⌝
→
Φ
k
x
→
Ψ
k
x
)
∧
([
∗
list
]
k
↦
x
∈
l
,
Φ
k
x
)
⊢
[
∗
list
]
k
↦
x
∈
l
,
Ψ
k
x
.
Proof
.
rewrite
always
_and_sep_l
.
do
2
setoid_rewrite
always
_forall
.
setoid_rewrite
always
_impl
;
setoid_rewrite
always
_pure
.
rewrite
persistently
_and_sep_l
.
do
2
setoid_rewrite
persistently
_forall
.
setoid_rewrite
persistently
_impl
;
setoid_rewrite
persistently
_pure
.
rewrite
-
big_sepL_forall
-
big_sepL_sepL
.
apply
big_sepL_mono
;
auto
=>
k
x
?.
by
rewrite
-
always_wand_impl
always
_elim
wand_elim_l
.
by
rewrite
-
persistently_wand_impl
persistently
_elim
wand_elim_l
.
Qed
.
Global
Instance
big_sepL_nil_persistent
Φ
:
...
...
@@ -307,11 +307,11 @@ Section gmap.
▷
^
n
([
∗
map
]
k
↦
x
∈
m
,
Φ
k
x
)
⊣
⊢
([
∗
map
]
k
↦
x
∈
m
,
▷
^
n
Φ
k
x
).
Proof
.
apply
(
big_opM_commute
_
).
Qed
.
Lemma
big_sepM_
always
Φ
m
:
Lemma
big_sepM_
persistently
Φ
m
:
(
□
[
∗
map
]
k
↦
x
∈
m
,
Φ
k
x
)
⊣
⊢
([
∗
map
]
k
↦
x
∈
m
,
□
Φ
k
x
).
Proof
.
apply
(
big_opM_commute
_
).
Qed
.
Lemma
big_sepM_
always
_if
p
Φ
m
:
Lemma
big_sepM_
persistently
_if
p
Φ
m
:
□
?p
([
∗
map
]
k
↦
x
∈
m
,
Φ
k
x
)
⊣
⊢
([
∗
map
]
k
↦
x
∈
m
,
□
?p
Φ
k
x
).
Proof
.
apply
(
big_opM_commute
_
).
Qed
.
...
...
@@ -323,7 +323,7 @@ Section gmap.
{
apply
forall_intro
=>
k
;
apply
forall_intro
=>
x
.
apply
impl_intro_l
,
pure_elim_l
=>
?
;
by
apply
big_sepM_lookup
.
}
induction
m
as
[|
i
x
m
?
IH
]
using
map_ind
;
[
rewrite
?big_sepM_empty
;
auto
|].
rewrite
big_sepM_insert
//
-
always
_and_sep_l
.
apply
and_intro
.
rewrite
big_sepM_insert
//
-
persistently
_and_sep_l
.
apply
and_intro
.
-
rewrite
(
forall_elim
i
)
(
forall_elim
x
)
lookup_insert
.
by
rewrite
pure_True
//
True_impl
.
-
rewrite
-
IH
.
apply
forall_mono
=>
k
;
apply
forall_mono
=>
y
.
...
...
@@ -336,10 +336,10 @@ Section gmap.
□
(
∀
k
x
,
⌜
m
!!
k
=
Some
x
⌝
→
Φ
k
x
→
Ψ
k
x
)
∧
([
∗
map
]
k
↦
x
∈
m
,
Φ
k
x
)
⊢
[
∗
map
]
k
↦
x
∈
m
,
Ψ
k
x
.
Proof
.
rewrite
always
_and_sep_l
.
do
2
setoid_rewrite
always
_forall
.
setoid_rewrite
always
_impl
;
setoid_rewrite
always
_pure
.
rewrite
persistently
_and_sep_l
.
do
2
setoid_rewrite
persistently
_forall
.
setoid_rewrite
persistently
_impl
;
setoid_rewrite
persistently
_pure
.
rewrite
-
big_sepM_forall
-
big_sepM_sepM
.
apply
big_sepM_mono
;
auto
=>
k
x
?.
by
rewrite
-
always_wand_impl
always
_elim
wand_elim_l
.
by
rewrite
-
persistently_wand_impl
persistently
_elim
wand_elim_l
.
Qed
.
Global
Instance
big_sepM_empty_persistent
Φ
:
...
...
@@ -460,10 +460,10 @@ Section gset.
▷
^
n
([
∗
set
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
set
]
y
∈
X
,
▷
^
n
Φ
y
).
Proof
.
apply
(
big_opS_commute
_
).
Qed
.
Lemma
big_sepS_
always
Φ
X
:
□
([
∗
set
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
set
]
y
∈
X
,
□
Φ
y
).
Lemma
big_sepS_
persistently
Φ
X
:
□
([
∗
set
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
set
]
y
∈
X
,
□
Φ
y
).
Proof
.
apply
(
big_opS_commute
_
).
Qed
.
Lemma
big_sepS_
always
_if
q
Φ
X
:
Lemma
big_sepS_
persistently
_if
q
Φ
X
:
□
?q
([
∗
set
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
set
]
y
∈
X
,
□
?q
Φ
y
).
Proof
.
apply
(
big_opS_commute
_
).
Qed
.
...
...
@@ -475,7 +475,7 @@ Section gset.
apply
impl_intro_l
,
pure_elim_l
=>
?
;
by
apply
big_sepS_elem_of
.
}
induction
X
as
[|
x
X
?
IH
]
using
collection_ind_L
.
{
rewrite
big_sepS_empty
;
auto
.
}
rewrite
big_sepS_insert
//
-
always
_and_sep_l
.
apply
and_intro
.
rewrite
big_sepS_insert
//
-
persistently
_and_sep_l
.
apply
and_intro
.
-
by
rewrite
(
forall_elim
x
)
pure_True
?True_impl
;
last
set_solver
.
-
rewrite
-
IH
.
apply
forall_mono
=>
y
.
apply
impl_intro_l
,
pure_elim_l
=>
?.
by
rewrite
pure_True
?True_impl
;
last
set_solver
.
...
...
@@ -484,10 +484,10 @@ Section gset.
Lemma
big_sepS_impl
Φ
Ψ
X
:
□
(
∀
x
,
⌜
x
∈
X
⌝
→
Φ
x
→
Ψ
x
)
∧
([
∗
set
]
x
∈
X
,
Φ
x
)
⊢
[
∗
set
]
x
∈
X
,
Ψ
x
.
Proof
.
rewrite
always_and_sep_l
always
_forall
.
setoid_rewrite
always
_impl
;
setoid_rewrite
always
_pure
.
rewrite
persistently_and_sep_l
persistently
_forall
.
setoid_rewrite
persistently
_impl
;
setoid_rewrite
persistently
_pure
.
rewrite
-
big_sepS_forall
-
big_sepS_sepS
.
apply
big_sepS_mono
;
auto
=>
x
?.
by
rewrite
-
always_wand_impl
always
_elim
wand_elim_l
.
by
rewrite
-
persistently_wand_impl
persistently
_elim
wand_elim_l
.
Qed
.
Global
Instance
big_sepS_empty_persistent
Φ
:
Persistent
([
∗
set
]
x
∈
∅
,
Φ
x
).
...
...
@@ -571,10 +571,10 @@ Section gmultiset.
▷
^
n
([
∗
mset
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
mset
]
y
∈
X
,
▷
^
n
Φ
y
).
Proof
.
apply
(
big_opMS_commute
_
).
Qed
.
Lemma
big_sepMS_
always
Φ
X
:
□
([
∗
mset
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
mset
]
y
∈
X
,
□
Φ
y
).
Lemma
big_sepMS_
persistently
Φ
X
:
□
([
∗
mset
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
mset
]
y
∈
X
,
□
Φ
y
).
Proof
.
apply
(
big_opMS_commute
_
).
Qed
.
Lemma
big_sepMS_
always
_if
q
Φ
X
:
Lemma
big_sepMS_
persistently
_if
q
Φ
X
:
□
?q
([
∗
mset
]
y
∈
X
,
Φ
y
)
⊣
⊢
([
∗
mset
]
y
∈
X
,
□
?q
Φ
y
).
Proof
.
apply
(
big_opMS_commute
_
).
Qed
.
...
...
theories/base_logic/derived.v
View file @
b4567fbd
This diff is collapsed.
Click to expand it.
theories/base_logic/lib/fractional.v
View file @
b4567fbd
...
...
@@ -51,7 +51,7 @@ Section fractional.
(** Fractional and logical connectives *)
Global
Instance
persistent_fractional
P
:
Persistent
P
→
Fractional
(
λ
_
,
P
).
Proof
.
intros
HP
q
q'
.
by
apply
uPred
.
always
_sep_dup
.
Qed
.
Proof
.
intros
HP
q
q'
.
by
apply
uPred
.
persistently
_sep_dup
.
Qed
.
Global
Instance
fractional_sep
Φ
Ψ
:
Fractional
Φ
→
Fractional
Ψ
→
Fractional
(
λ
q
,
Φ
q
∗
Ψ
q
)%
I
.
...
...
@@ -134,7 +134,7 @@ Section fractional.
AsFractional
P
Φ
(
q1
+
q2
)
→
AsFractional
P1
Φ
q1
→
AsFractional
P2
Φ
q2
→
IntoAnd
p
P
P1
P2
.
Proof
.
(* TODO: We need a better way to handle this boolean here;
always
(* TODO: We need a better way to handle this boolean here;
persistently
applying mk_into_and_sep (which only works after introducing all
assumptions) is rather annoying.
Ideally, it'd not even be possible to make the mistake that
...
...
@@ -148,7 +148,7 @@ Section fractional.
Proof
.
intros
.
apply
mk_into_and_sep
.
rewrite
[
P
]
fractional_half
//.
Qed
.
(* The instance [frame_fractional] can be tried at all the nodes of
the proof search. The proof search then fails almost
always
on
the proof search. The proof search then fails almost
persistently
on
[AsFractional R Φ r], but the slowdown is still noticeable. For
that reason, we factorize the three instances that could have been
defined for that purpose into one. *)
...
...
@@ -179,6 +179,6 @@ Section fractional.
-
rewrite
fractional
=><-<-.
by
rewrite
assoc
.
-
rewrite
fractional
=><-<-=>
_
.
by
rewrite
(
comm
_
Q
(
Φ
q0
))
!
assoc
(
comm
_
(
Φ
_
)).
-
move
=>-[->
_
]->.
by
rewrite
uPred
.
always
_if_elim
-
fractional
Qp_div_2
.
-
move
=>-[->
_
]->.
by
rewrite
uPred
.
persistently
_if_elim
-
fractional
Qp_div_2
.
Qed
.
End
fractional
.
theories/base_logic/lib/iprop.v
View file @
b4567fbd
...
...
@@ -83,7 +83,7 @@ Class subG (Σ1 Σ2 : gFunctors) := in_subG i : { j | Σ1 i = Σ2 j }.
(** Avoid trigger happy type class search: this line ensures that type class
search is only triggered if the arguments of [subG] do not contain evars. Since
instance search for [subG] is restrained, instances should
always
have [subG] as
instance search for [subG] is restrained, instances should
persistently
have [subG] as
their first parameter to avoid loops. For example, the instances [subG_authΣ]
and [auth_discrete] otherwise create a cycle that pops up arbitrarily. *)
Hint
Mode
subG
+
+
:
typeclass_instances
.
...
...
theories/base_logic/lib/own.v
View file @
b4567fbd
...
...
@@ -102,7 +102,7 @@ Proof. apply wand_intro_r. by rewrite -own_op own_valid. Qed.
Lemma
own_valid_3
γ
a1
a2
a3
:
own
γ
a1
-
∗
own
γ
a2
-
∗
own
γ
a3
-
∗
✓
(
a1
⋅
a2
⋅
a3
).
Proof
.
do
2
apply
wand_intro_r
.
by
rewrite
-!
own_op
own_valid
.
Qed
.
Lemma
own_valid_r
γ
a
:
own
γ
a
⊢
own
γ
a
∗
✓
a
.
Proof
.
apply
:
uPred
.
always
_entails_r
.
apply
own_valid
.
Qed
.
Proof
.
apply
:
uPred
.
persistently
_entails_r
.
apply
own_valid
.
Qed
.
Lemma
own_valid_l
γ
a
:
own
γ
a
⊢
✓
a
∗
own
γ
a
.
Proof
.
by
rewrite
comm
-
own_valid_r
.
Qed
.
...
...
theories/base_logic/lib/viewshifts.v
View file @
b4567fbd
...
...
@@ -81,5 +81,5 @@ Lemma vs_alloc N P : ▷ P ={↑N}=> inv N P.
Proof
.
iIntros
"!# HP"
.
by
iApply
inv_alloc
.
Qed
.
Lemma
wand_fupd_alt
E1
E2
P
Q
:
(
P
={
E1
,
E2
}=
∗
Q
)
⊣
⊢
∃
R
,
R
∗
(
P
∗
R
={
E1
,
E2
}=>
Q
).
Proof
.
rewrite
uPred
.
wand_alt
.
by
setoid_rewrite
<-
uPred
.
always
_wand_impl
.
Qed
.
Proof
.
rewrite
uPred
.
wand_alt
.
by
setoid_rewrite
<-
uPred
.
persistently
_wand_impl
.
Qed
.
End
vs
.
theories/base_logic/primitive.v
View file @
b4567fbd
...
...
@@ -97,16 +97,16 @@ Definition uPred_wand {M} := unseal uPred_wand_aux M.
Definition
uPred_wand_eq
:
@
uPred_wand
=
@
uPred_wand_def
:
=
seal_eq
uPred_wand_aux
.
Program
Definition
uPred_
always
_def
{
M
}
(
P
:
uPred
M
)
:
uPred
M
:
=
Program
Definition
uPred_
persistently
_def
{
M
}
(
P
:
uPred
M
)
:
uPred
M
:
=
{|
uPred_holds
n
x
:
=
P
n
(
core
x
)
|}.
Next
Obligation
.
intros
M
;
naive_solver
eauto
using
uPred_mono
,
@
cmra_core_monoN
.
Qed
.
Next
Obligation
.
naive_solver
eauto
using
uPred_closed
,
@
cmra_core_validN
.
Qed
.
Definition
uPred_
always
_aux
:
seal
(@
uPred_
always
_def
).
by
eexists
.
Qed
.
Definition
uPred_
always
{
M
}
:
=
unseal
uPred_
always
_aux
M
.
Definition
uPred_
always
_eq
:
@
uPred_
always
=
@
uPred_always
_def
:
=
seal_eq
uPred_
always
_aux
.
Definition
uPred_
persistently
_aux
:
seal
(@
uPred_
persistently
_def
).
by
eexists
.
Qed
.
Definition
uPred_
persistently
{
M
}
:
=
unseal
uPred_
persistently
_aux
M
.
Definition
uPred_
persistently
_eq
:
@
uPred_
persistently
=
@
uPred_persistently
_def
:
=
seal_eq
uPred_
persistently
_aux
.
Program
Definition
uPred_later_def
{
M
}
(
P
:
uPred
M
)
:
uPred
M
:
=
{|
uPred_holds
n
x
:
=
match
n
return
_
with
0
=>
True
|
S
n'
=>
P
n'
x
end
|}.
...
...
@@ -176,7 +176,7 @@ Notation "∀ x .. y , P" :=
Notation
"∃ x .. y , P"
:
=
(
uPred_exist
(
λ
x
,
..
(
uPred_exist
(
λ
y
,
P
))
..)%
I
)
(
at
level
200
,
x
binder
,
y
binder
,
right
associativity
)
:
uPred_scope
.
Notation
"□ P"
:
=
(
uPred_
always
P
)
Notation
"□ P"
:
=
(
uPred_
persistently
P
)
(
at
level
20
,
right
associativity
)
:
uPred_scope
.
Notation
"▷ P"
:
=
(
uPred_later
P
)
(
at
level
20
,
right
associativity
)
:
uPred_scope
.
...
...
@@ -198,7 +198,7 @@ Notation "P -∗ Q" := (P ⊢ Q)
Module
uPred
.
Definition
unseal_eqs
:
=
(
uPred_pure_eq
,
uPred_and_eq
,
uPred_or_eq
,
uPred_impl_eq
,
uPred_forall_eq
,
uPred_exist_eq
,
uPred_internal_eq_eq
,
uPred_sep_eq
,
uPred_wand_eq
,
uPred_
always
_eq
,
uPred_exist_eq
,
uPred_internal_eq_eq
,
uPred_sep_eq
,
uPred_wand_eq
,
uPred_
persistently
_eq
,
uPred_later_eq
,
uPred_ownM_eq
,
uPred_cmra_valid_eq
,
uPred_bupd_eq
).
Ltac
unseal
:
=
rewrite
!
unseal_eqs
/=.
...
...
@@ -295,13 +295,13 @@ Proof.
Qed
.
Global
Instance
later_proper'
:
Proper
((
⊣
⊢
)
==>
(
⊣
⊢
))
(@
uPred_later
M
)
:
=
ne_proper
_
.
Global
Instance
always
_ne
:
NonExpansive
(@
uPred_
always
M
).
Global
Instance
persistently
_ne
:
NonExpansive
(@
uPred_
persistently
M
).
Proof
.
intros
n
P1
P2
HP
.
unseal
;
split
=>
n'
x
;
split
;
apply
HP
;
eauto
using
@
cmra_core_validN
.
Qed
.
Global
Instance
always
_proper
:
Proper
((
⊣
⊢
)
==>
(
⊣
⊢
))
(@
uPred_
always
M
)
:
=
ne_proper
_
.
Global
Instance
persistently
_proper
:
Proper
((
⊣
⊢
)
==>
(
⊣
⊢
))
(@
uPred_
persistently
M
)
:
=
ne_proper
_
.
Global
Instance
ownM_ne
:
NonExpansive
(@
uPred_ownM
M
).
Proof
.
intros
n
a
b
Ha
.
...
...
@@ -422,22 +422,22 @@ Proof.
Qed
.
(* Always *)
Lemma
always
_mono
P
Q
:
(
P
⊢
Q
)
→
□
P
⊢
□
Q
.
Lemma
persistently
_mono
P
Q
:
(
P
⊢
Q
)
→
□
P
⊢
□
Q
.
Proof
.
intros
HP
;
unseal
;
split
=>
n
x
?
/=.
by
apply
HP
,
cmra_core_validN
.
Qed
.
Lemma
always
_elim
P
:
□
P
⊢
P
.
Lemma
persistently
_elim
P
:
□
P
⊢
P
.
Proof
.
unseal
;
split
=>
n
x
?
/=.
eauto
using
uPred_mono
,
@
cmra_included_core
,
cmra_included_includedN
.
Qed
.
Lemma
always
_idemp_2
P
:
□
P
⊢
□
□
P
.
Lemma
persistently
_idemp_2
P
:
□
P
⊢
□
□
P
.
Proof
.
unseal
;
split
=>
n
x
??
/=.
by
rewrite
cmra_core_idemp
.
Qed
.
Lemma
always
_forall_2
{
A
}
(
Ψ
:
A
→
uPred
M
)
:
(
∀
a
,
□
Ψ
a
)
⊢
(
□
∀
a
,
Ψ
a
).
Lemma
persistently
_forall_2
{
A
}
(
Ψ
:
A
→
uPred
M
)
:
(
∀
a
,
□
Ψ
a
)
⊢
(
□
∀
a
,
Ψ
a
).
Proof
.
by
unseal
.
Qed
.
Lemma
always
_exist_1
{
A
}
(
Ψ
:
A
→
uPred
M
)
:
(
□
∃
a
,
Ψ
a
)
⊢
(
∃
a
,
□
Ψ
a
).
Lemma
persistently
_exist_1
{
A
}
(
Ψ
:
A
→
uPred
M
)
:
(
□
∃
a
,
Ψ
a
)
⊢
(
∃
a
,
□
Ψ
a
).
Proof
.
by
unseal
.
Qed
.
Lemma
always
_and_sep_l_1
P
Q
:
□
P
∧
Q
⊢
□
P
∗
Q
.
Lemma
persistently
_and_sep_l_1
P
Q
:
□
P
∧
Q
⊢
□
P
∗
Q
.
Proof
.
unseal
;
split
=>
n
x
?
[??]
;
exists
(
core
x
),
x
;
simpl
in
*.
by
rewrite
cmra_core_l
cmra_core_idemp
.
...
...
@@ -475,7 +475,7 @@ Proof.
intros
[|
n'
]
x'
????
;
[|
done
].
eauto
using
uPred_closed
,
uPred_mono
,
cmra_included_includedN
.
Qed
.
Lemma
always
_later
P
:
□
▷
P
⊣
⊢
▷
□
P
.
Lemma
persistently
_later
P
:
□
▷
P
⊣
⊢
▷
□
P
.
Proof
.
by
unseal
.
Qed
.
(* Own *)
...
...
@@ -489,7 +489,7 @@ Proof.
by
rewrite
(
assoc
op
_
z1
)
-(
comm
op
z1
)
(
assoc
op
z1
)
-(
assoc
op
_
a2
)
(
comm
op
z1
)
-
Hy1
-
Hy2
.
Qed
.
Lemma
always
_ownM_core
(
a
:
M
)
:
uPred_ownM
a
⊢
□
uPred_ownM
(
core
a
).
Lemma
persistently
_ownM_core
(
a
:
M
)
:
uPred_ownM
a
⊢
□
uPred_ownM
(
core
a
).
Proof
.
split
=>
n
x
/=
;
unseal
;
intros
Hx
.
simpl
.
by
apply
cmra_core_monoN
.
Qed
.
...
...
@@ -512,7 +512,7 @@ Lemma cmra_valid_intro {A : cmraT} (a : A) : ✓ a → uPred_valid (M:=M) (✓ a
Proof
.
unseal
=>
?
;
split
=>
n
x
?
_
/=
;
by
apply
cmra_valid_validN
.
Qed
.
Lemma
cmra_valid_elim
{
A
:
cmraT
}
(
a
:
A
)
:
¬
✓
{
0
}
a
→
✓
a
⊢
False
.
Proof
.
unseal
=>
Ha
;
split
=>
n
x
??
;
apply
Ha
,
cmra_validN_le
with
n
;
auto
.
Qed
.
Lemma
always
_cmra_valid_1
{
A
:
cmraT
}
(
a
:
A
)
:
✓
a
⊢
□
✓
a
.
Lemma
persistently
_cmra_valid_1
{
A
:
cmraT
}
(
a
:
A
)
:
✓
a
⊢
□
✓
a
.
Proof
.
by
unseal
.
Qed
.
Lemma
cmra_valid_weaken
{
A
:
cmraT
}
(
a
b
:
A
)
:
✓
(
a
⋅
b
)
⊢
✓
a
.
Proof
.
unseal
;
split
=>
n
x
_;
apply
cmra_validN_op_l
.
Qed
.
...
...
theories/program_logic/hoare.v
View file @
b4567fbd
...
...
@@ -37,7 +37,7 @@ Global Instance ht_proper E :
Proof
.
solve_proper
.
Qed
.
Lemma
ht_mono
E
P
P'
Φ
Φ
'
e
:
(
P
⊢
P'
)
→
(
∀
v
,
Φ
'
v
⊢
Φ
v
)
→
{{
P'
}}
e
@
E
{{
Φ
'
}}
⊢
{{
P
}}
e
@
E
{{
Φ
}}.
Proof
.
by
intros
;
apply
always
_mono
,
wand_mono
,
wp_mono
.
Qed
.
Proof
.
by
intros
;
apply
persistently
_mono
,
wand_mono
,
wp_mono
.
Qed
.
Global
Instance
ht_mono'
E
:
Proper
(
flip
(
⊢
)
==>
eq
==>
pointwise_relation
_
(
⊢
)
==>
(
⊢
))
(
ht
E
).
Proof
.
solve_proper
.
Qed
.
...
...
theories/program_logic/weakestpre.v
View file @
b4567fbd
...
...
@@ -283,7 +283,7 @@ Section proofmode_classes.
ElimModal
(|={
E
}=>
P
)
P
(
WP
e
@
E
{{
Φ
}})
(
WP
e
@
E
{{
Φ
}}).
Proof
.
by
rewrite
/
ElimModal
fupd_frame_r
wand_elim_r
fupd_wp
.
Qed
.
(* lower precedence, if possible, it should
always
pick elim_upd_fupd_wp *)
(* lower precedence, if possible, it should
persistently
pick elim_upd_fupd_wp *)
Global
Instance
elim_modal_fupd_wp_atomic
E1
E2
e
P
Φ
:
atomic
e
→
ElimModal
(|={
E1
,
E2
}=>
P
)
P
...
...
theories/proofmode/class_instances.v
View file @
b4567fbd
...
...
@@ -11,17 +11,17 @@ Implicit Types P Q R : uPred M.
(* FromAssumption *)
Global
Instance
from_assumption_exact
p
P
:
FromAssumption
p
P
P
|
0
.
Proof
.
destruct
p
;
by
rewrite
/
FromAssumption
/=
?
always
_elim
.
Qed
.
Proof
.
destruct
p
;
by
rewrite
/
FromAssumption
/=
?
persistently
_elim
.
Qed
.
Global
Instance
from_assumption_False
p
P
:
FromAssumption
p
False
P
|
1
.
Proof
.
destruct
p
;
rewrite
/
FromAssumption
/=
?
always
_pure
;
apply
False_elim
.
Qed
.
Proof
.
destruct
p
;
rewrite
/
FromAssumption
/=
?
persistently
_pure
;
apply
False_elim
.
Qed
.
Global
Instance
from_assumption_
always
_r
P
Q
:
Global
Instance
from_assumption_
persistently
_r
P
Q
:
FromAssumption
true
P
Q
→
FromAssumption
true
P
(
□
Q
).
Proof
.
rewrite
/
FromAssumption
=><-.
by
rewrite
always_always
.
Qed
.
Proof
.
rewrite
/
FromAssumption
=><-.
by
rewrite
persistently_persistently
.
Qed
.
Global
Instance
from_assumption_
always
_l
p
P
Q
:
Global
Instance
from_assumption_
persistently
_l
p
P
Q
:
FromAssumption
p
P
Q
→
FromAssumption
p
(
□
P
)
Q
.
Proof
.
rewrite
/
FromAssumption
=><-.
by
rewrite
always
_elim
.
Qed
.
Proof
.
rewrite
/
FromAssumption
=><-.
by
rewrite
persistently
_elim
.
Qed
.
Global
Instance
from_assumption_later
p
P
Q
:
FromAssumption
p
P
Q
→
FromAssumption
p
P
(
▷
Q
)%
I
.
Proof
.
rewrite
/
FromAssumption
=>->.
apply
later_intro
.
Qed
.
...
...
@@ -48,8 +48,8 @@ Global Instance into_pure_cmra_valid `{CmraDiscrete A} (a : A) :
@
IntoPure
M
(
✓
a
)
(
✓
a
).
Proof
.
by
rewrite
/
IntoPure
discrete_valid
.
Qed
.
Global
Instance
into_pure_
always
P
φ
:
IntoPure
P
φ
→
IntoPure
(
□
P
)
φ
.
Proof
.
rewrite
/
IntoPure
=>
->.
by
rewrite
always
_pure
.
Qed
.
Global
Instance
into_pure_
persistently
P
φ
:
IntoPure
P
φ
→
IntoPure
(
□
P
)
φ
.
Proof
.
rewrite
/
IntoPure
=>
->.
by
rewrite
persistently
_pure
.
Qed
.
Global
Instance
into_pure_pure_and
(
φ
1
φ
2
:
Prop
)
P1
P2
:
IntoPure
P1
φ
1
→
IntoPure
P2
φ
2
→
IntoPure
(
P1
∧
P2
)
(
φ
1
∧
φ
2
).
...
...
@@ -66,7 +66,7 @@ Proof. rewrite /FromPure /IntoPure pure_impl. by intros -> ->. Qed.
Global
Instance
into_pure_pure_wand
(
φ
1
φ
2
:
Prop
)
P1
P2
:
FromPure
P1
φ
1
→
IntoPure
P2
φ
2
→
IntoPure
(
P1
-
∗
P2
)
(
φ
1
→
φ
2
).
Proof
.
rewrite
/
FromPure
/
IntoPure
pure_impl
always
_impl_wand
.
by
intros
->
->.
rewrite
/
FromPure
/
IntoPure
pure_impl
persistently
_impl_wand
.
by
intros
->
->.
Qed
.
Global
Instance
into_pure_exist
{
X
:
Type
}
(
Φ
:
X
→
uPred
M
)
(
φ
:
X
→
Prop
)
:
...
...
@@ -97,8 +97,8 @@ Proof.
rewrite
-
cmra_valid_intro
//.
auto
with
I
.
Qed
.
Global
Instance
from_pure_
always
P
φ
:
FromPure
P
φ
→
FromPure
(
□
P
)
φ
.
Proof
.
rewrite
/
FromPure
=>
<-.
by
rewrite
always
_pure
.
Qed
.
Global
Instance
from_pure_
persistently
P
φ
:
FromPure
P
φ
→
FromPure
(
□
P
)
φ
.
Proof
.
rewrite
/
FromPure
=>
<-.
by
rewrite
persistently
_pure
.
Qed
.
Global
Instance
from_pure_later
P
φ
:
FromPure
P
φ
→
FromPure
(
▷
P
)
φ
.
Proof
.
rewrite
/
FromPure
=>
->.
apply
later_intro
.
Qed
.
Global
Instance
from_pure_laterN
n
P
φ
:
FromPure
P
φ
→
FromPure
(
▷
^
n
P
)
φ
.
...
...
@@ -113,7 +113,7 @@ Global Instance from_pure_pure_and (φ1 φ2 : Prop) P1 P2 :
Proof
.
rewrite
/
FromPure
pure_and
.
by
intros
->
->.
Qed
.
Global
Instance
from_pure_pure_sep
(
φ
1
φ
2
:
Prop
)
P1
P2
:
FromPure
P1
φ
1
→
FromPure
P2
φ
2
→
FromPure
(
P1
∗
P2
)
(
φ
1
∧
φ
2
).
Proof
.
rewrite
/
FromPure
pure_and
always
_and_sep_l
.
by
intros
->
->.
Qed
.
Proof
.
rewrite
/
FromPure
pure_and
persistently
_and_sep_l
.
by
intros
->
->.
Qed
.
Global
Instance
from_pure_pure_or
(
φ
1
φ
2
:
Prop
)
P1
P2
:
FromPure
P1
φ
1
→
FromPure
P2
φ
2
→
FromPure
(
P1
∨
P2
)
(
φ
1
∨
φ
2
).
Proof
.
rewrite
/
FromPure
pure_or
.
by
intros
->
->.
Qed
.
...
...
@@ -123,7 +123,7 @@ Proof. rewrite /FromPure /IntoPure pure_impl. by intros -> ->. Qed.
Global
Instance
from_pure_pure_wand
(
φ
1
φ
2
:
Prop
)
P1
P2
:
IntoPure
P1
φ
1
→
FromPure
P2
φ
2
→
FromPure
(
P1
-
∗
P2
)
(
φ
1
→
φ
2
).
Proof
.
rewrite
/
FromPure
/
IntoPure
pure_impl
always
_impl_wand
.
by
intros
->
->.
rewrite
/
FromPure
/
IntoPure
pure_impl
persistently
_impl_wand
.
by
intros
->
->.
Qed
.
Global
Instance
from_pure_exist
{
X
:
Type
}
(
Φ
:
X
→
uPred
M
)
(
φ
:
X
→
Prop
)
:
...
...
@@ -140,10 +140,10 @@ Proof.
Qed
.
(* IntoPersistent *)
Global
Instance
into_persistent_
always
_trans
p
P
Q
:
Global
Instance
into_persistent_
persistently
_trans
p
P
Q
:
IntoPersistent
true
P
Q
→
IntoPersistent
p
(
□
P
)
Q
|
0
.
Proof
.
rewrite
/
IntoPersistent
/==>
->.
by
rewrite
always_if_always
.
Qed
.
Global
Instance
into_persistent_
always
P
:
IntoPersistent
true
P
P
|
1
.
Proof
.
rewrite
/
IntoPersistent
/==>
->.
by
rewrite
persistently_if_persistently
.
Qed
.
Global
Instance
into_persistent_
persistently
P
:
IntoPersistent
true
P
P
|
1
.
Proof
.
done
.
Qed
.
Global
Instance
into_persistent_persistent
P
:
Persistent
P
→
IntoPersistent
false
P
P
|
100
.
...
...
@@ -258,9 +258,9 @@ Global Instance from_later_sep n P1 P2 Q1 Q2 :
FromLaterN
n
P1
Q1
→
FromLaterN
n
P2
Q2
→
FromLaterN
n
(
P1
∗
P2
)
(
Q1
∗
Q2
).
Proof
.
intros
??
;
red
.
by
rewrite
laterN_sep
;
apply
sep_mono
.
Qed
.
Global
Instance
from_later_
always
n
P
Q
:
Global
Instance
from_later_
persistently
n
P
Q
:
FromLaterN
n
P
Q
→
FromLaterN
n
(
□
P
)
(
□
Q
).
Proof
.
by
rewrite
/
FromLaterN
-
always
_laterN
=>
->.
Qed
.
Proof
.
by
rewrite
/
FromLaterN
-
persistently
_laterN
=>
->.
Qed
.
Global
Instance
from_later_forall
{
A
}
n
(
Φ
Ψ
:
A
→
uPred
M
)
:
(
∀
x
,
FromLaterN
n
(
Φ
x
)
(
Ψ
x
))
→
FromLaterN
n
(
∀
x
,
Φ
x
)
(
∀
x
,
Ψ
x
).
...
...
@@ -278,19 +278,19 @@ Global Instance wand_weaken_later p P Q P' Q' :
WandWeaken
p
P
Q
P'
Q'
→
WandWeaken'
p
P
Q
(
▷
P'
)
(
▷
Q'
).
Proof
.
rewrite
/
WandWeaken'
/
WandWeaken
=>
->.
by
rewrite
always
_if_later
-
later_wand
-
later_intro
.
by
rewrite
persistently
_if_later
-
later_wand
-
later_intro
.
Qed
.
Global
Instance
wand_weaken_laterN
p
n
P
Q
P'
Q'
:
WandWeaken
p
P
Q
P'
Q'
→
WandWeaken'
p
P
Q
(
▷
^
n
P'
)
(
▷
^
n
Q'
).
Proof
.
rewrite
/
WandWeaken'
/
WandWeaken
=>
->.
by
rewrite
always
_if_laterN
-
laterN_wand
-
laterN_intro
.
by
rewrite
persistently
_if_laterN
-
laterN_wand
-
laterN_intro
.
Qed
.
Global
Instance
bupd_weaken_laterN
p
P
Q
P'
Q'
:
WandWeaken
false
P
Q
P'
Q'
→
WandWeaken'
p
P
Q
(|==>
P'
)
(|==>
Q'
).
Proof
.
rewrite
/
WandWeaken'
/
WandWeaken
=>
->.
apply
wand_intro_l
.
by
rewrite
always
_if_elim
bupd_wand_r
.
apply
wand_intro_l
.
by
rewrite
persistently
_if_elim
bupd_wand_r
.
Qed
.
Global
Instance
into_wand_wand
p
P
P'
Q
Q'
:
...
...
@@ -310,16 +310,16 @@ Proof. rewrite /WandWeaken /IntoWand=> <-. apply and_elim_r', impl_wand. Qed.
Global
Instance
into_wand_forall
{
A
}
p
(
Φ
:
A
→
uPred
M
)
P
Q
x
:
IntoWand
p
(
Φ
x
)
P
Q
→
IntoWand
p
(
∀
x
,
Φ
x
)
P
Q
.
Proof
.
rewrite
/
IntoWand
=>
<-.
apply
forall_elim
.
Qed
.
Global
Instance
into_wand_
always
p
R
P
Q
:
Global
Instance
into_wand_
persistently
p
R
P
Q
:
IntoWand
p
R
P
Q
→
IntoWand
p
(
□
R
)
P
Q
.
Proof
.
rewrite
/
IntoWand
=>
->.
apply
always
_elim
.
Qed
.
Proof
.
rewrite
/
IntoWand
=>
->.
apply
persistently
_elim
.
Qed
.
Global
Instance
into_wand_later
p
R
P
Q
:
IntoWand
p
R
P
Q
→
IntoWand
p
(
▷
R
)
(
▷
P
)
(
▷
Q
).
Proof
.
rewrite
/
IntoWand
=>
->.
by
rewrite
always
_if_later
-
later_wand
.
Qed
.
Proof
.
rewrite
/
IntoWand
=>
->.
by
rewrite
persistently
_if_later
-
later_wand
.
Qed
.
Global
Instance
into_wand_laterN
p
n
R
P
Q
:
IntoWand
p
R
P
Q
→
IntoWand
p
(
▷
^
n
R
)
(
▷
^
n
P
)
(
▷
^
n
Q
).
Proof
.
rewrite
/
IntoWand
=>
->.
by
rewrite
always
_if_laterN
-
laterN_wand
.
Qed
.
Proof
.
rewrite
/
IntoWand
=>
->.
by
rewrite
persistently
_if_laterN
-
laterN_wand
.
Qed
.
Global
Instance
into_wand_bupd
R
P
Q
:
IntoWand
false
R
P
Q
→
IntoWand
false
(|==>
R
)
(|==>
P
)
(|==>
Q
).
...
...
@@ -340,18 +340,18 @@ Global Instance from_and_sep P1 P2 : FromAnd false (P1 ∗ P2) P1 P2 | 100.
Proof
.
done
.
Qed
.
Global
Instance
from_and_sep_persistent_l
P1
P2
:
Persistent
P1
→
FromAnd
true
(
P1
∗
P2
)
P1
P2
|
9
.
Proof
.
intros
.
by
rewrite
/
FromAnd
always
_and_sep_l
.
Qed
.
Proof
.
intros
.
by
rewrite
/
FromAnd
persistently
_and_sep_l
.
Qed
.
Global
Instance
from_and_sep_persistent_r
P1
P2
:
Persistent
P2
→
FromAnd
true
(
P1
∗
P2
)
P1
P2
|
10
.
Proof
.
intros
.
by
rewrite
/
FromAnd
always
_and_sep_r
.
Qed
.
Proof
.
intros
.
by
rewrite
/
FromAnd
persistently
_and_sep_r
.
Qed
.
Global
Instance
from_and_pure
p
φ
ψ
:
@
FromAnd
M
p
⌜φ
∧
ψ⌝
⌜φ⌝
⌜ψ⌝
.
Proof
.
apply
mk_from_and_and
.
by
rewrite
pure_and
.
Qed
.
Global
Instance
from_and_
always
p
P
Q1
Q2
:
Global
Instance
from_and_
persistently
p
P
Q1
Q2
:
FromAnd
false
P
Q1
Q2
→
FromAnd
p
(
□
P
)
(
□
Q1
)
(
□
Q2
).
Proof
.
intros
.
apply
mk_from_and_and
.
by
rewrite
always
_and_sep_l'
-
always
_sep
-(
from_and
_
P
).
by
rewrite
persistently
_and_sep_l'
-
persistently
_sep
-(
from_and
_
P
).
Qed
.
Global
Instance
from_and_later
p
P
Q1
Q2
:
FromAnd
p
P
Q1
Q2
→
FromAnd
p
(
▷
P
)
(
▷
Q1
)
(
▷
Q2
).
...
...
@@ -387,7 +387,7 @@ Proof. by rewrite /FromAnd big_sepL_cons. Qed.
Global
Instance
from_and_big_sepL_cons_persistent
{
A
}
(
Φ
:
nat
→
A
→
uPred
M
)
x
l
:
Persistent
(
Φ
0
x
)
→
FromAnd
true
([
∗
list
]
k
↦
y
∈
x
::
l
,
Φ
k
y
)
(
Φ
0
x
)
([
∗
list
]
k
↦
y
∈
l
,
Φ
(
S
k
)
y
).
Proof
.
intros
.
by
rewrite
/
FromAnd
big_opL_cons
always
_and_sep_l
.
Qed
.
Proof
.
intros
.
by
rewrite
/
FromAnd
big_opL_cons
persistently
_and_sep_l
.
Qed
.
Global
Instance
from_and_big_sepL_app
{
A
}
(
Φ
:
nat
→
A
→
uPred
M
)
l1
l2
:
FromAnd
false
([
∗
list
]
k
↦
y
∈
l1
++
l2
,
Φ
k
y
)
...
...
@@ -397,7 +397,7 @@ Global Instance from_sep_big_sepL_app_persistent {A} (Φ : nat → A → uPred M
(
∀
k
y
,
Persistent
(
Φ
k
y
))
→
FromAnd
true
([
∗
list
]
k
↦
y
∈
l1
++
l2
,
Φ
k
y
)
([
∗
list
]
k
↦
y
∈
l1
,
Φ
k
y
)
([
∗
list
]
k
↦
y
∈
l2
,
Φ
(
length
l1
+
k
)
y
).
Proof
.
intros
.
by
rewrite
/
FromAnd
big_opL_app
always
_and_sep_l
.
Qed
.
Proof
.
intros
.
by
rewrite
/
FromAnd
big_opL_app
persistently
_and_sep_l
.
Qed
.
(* FromOp *)
(* TODO: Worst case there could be a lot of backtracking on these instances,
...
...
@@ -431,17 +431,17 @@ Global Instance into_and_and P Q : IntoAnd true (P ∧ Q) P Q.
Proof
.
done
.
Qed
.
Global
Instance
into_and_and_persistent_l
P
Q
:
Persistent
P
→
IntoAnd
false
(
P
∧
Q
)
P
Q
.
Proof
.
intros
;
by
rewrite
/
IntoAnd
/=
always
_and_sep_l
.
Qed
.
Proof
.
intros
;
by
rewrite
/
IntoAnd
/=
persistently
_and_sep_l
.
Qed
.
Global
Instance
into_and_and_persistent_r
P
Q
:
Persistent
Q
→
IntoAnd
false
(
P
∧
Q
)
P
Q
.
Proof
.
intros
;
by
rewrite
/
IntoAnd
/=
always
_and_sep_r
.
Qed
.
Proof
.
intros
;
by
rewrite
/
IntoAnd
/=
persistently
_and_sep_r
.
Qed
.
Global
Instance
into_and_pure
p
φ
ψ
:
@
IntoAnd
M
p
⌜φ
∧
ψ⌝
⌜φ⌝
⌜ψ⌝
.
Proof
.
apply
mk_into_and_sep
.
by
rewrite
pure_and
always
_and_sep_r
.
Qed
.
Global
Instance
into_and_
always
p
P
Q1
Q2
:
Proof
.
apply
mk_into_and_sep
.
by
rewrite
pure_and
persistently
_and_sep_r
.
Qed
.
Global
Instance
into_and_
persistently
p
P
Q1
Q2
:
IntoAnd
true
P
Q1
Q2
→
IntoAnd
p
(
□
<