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
Jonas Kastberg
iris
Commits
568f6b7a
Commit
568f6b7a
authored
Oct 21, 2016
by
Robbert Krebbers
Browse files
Use notation |==> instead of |=r=> for basic update modality.
parent
1b85d654
Changes
21
Hide whitespace changes
Inline
Side-by-side
algebra/double_negation.v
View file @
568f6b7a
...
...
@@ -23,8 +23,8 @@ Notation "P =n=★ Q" := (P -★ |=n=> Q)%I
(
at
level
99
,
Q
at
level
200
,
format
"P =n=★ Q"
)
:
uPred_scope
.
(* Our goal is to prove that:
(1) |=n=> has (nearly) all the properties of the |=
r
=> modality that are used in Iris
(2) If our meta-logic is classical, then |=n=> and |=
r
=> are equivalent
(1) |=n=> has (nearly) all the properties of the |==> modality that are used in Iris
(2) If our meta-logic is classical, then |=n=> and |==> are equivalent
*)
Section
bupd_nnupd
.
...
...
@@ -264,7 +264,7 @@ Qed.
direction from bupd to nnupd is similar to the proof of
nnupd_ownM_updateP *)
Lemma
bupd_nnupd
P
:
(|=
r
=>
P
)
⊢
|=
n
=>
P
.
Lemma
bupd_nnupd
P
:
(|==>
P
)
⊢
|=
n
=>
P
.
Proof
.
split
.
rewrite
/
uPred_nnupd
.
repeat
uPred
.
unseal
.
intros
n
x
?
Hbupd
a
.
red
;
rewrite
//=
=>
n'
yf
??.
...
...
@@ -282,7 +282,7 @@ Qed.
(* However, the other direction seems to need a classical axiom: *)
Section
classical
.
Context
(
not_all_not_ex
:
∀
(
P
:
M
→
Prop
),
¬
(
∀
n
:
M
,
¬
P
n
)
→
∃
n
:
M
,
P
n
).
Lemma
nnupd_bupd
P
:
(|=
n
=>
P
)
⊢
(|=
r
=>
P
).
Lemma
nnupd_bupd
P
:
(|=
n
=>
P
)
⊢
(|==>
P
).
Proof
.
rewrite
/
uPred_nnupd
.
split
.
uPred
.
unseal
;
red
;
rewrite
//=.
...
...
@@ -373,8 +373,8 @@ Qed.
(* Open question:
Do the basic properties of the |=
r
=> modality (bupd_intro, bupd_mono, rvs_trans, rvs_frame_r,
bupd_ownM_updateP, and adequacy) uniquely characterize |=
r
=>?
Do the basic properties of the |==> modality (bupd_intro, bupd_mono, rvs_trans, rvs_frame_r,
bupd_ownM_updateP, and adequacy) uniquely characterize |==>?
*)
End
bupd_nnupd
.
algebra/upred.v
View file @
568f6b7a
...
...
@@ -310,12 +310,12 @@ Notation "▷ P" := (uPred_later P)
(
at
level
20
,
right
associativity
)
:
uPred_scope
.
Infix
"≡"
:
=
uPred_eq
:
uPred_scope
.
Notation
"✓ x"
:
=
(
uPred_cmra_valid
x
)
(
at
level
20
)
:
uPred_scope
.
Notation
"|=
r
=> Q"
:
=
(
uPred_bupd
Q
)
(
at
level
99
,
Q
at
level
200
,
format
"|=
r
=> Q"
)
:
uPred_scope
.
Notation
"P =
r=>
Q"
:
=
(
P
⊢
|=
r
=>
Q
)
Notation
"|==> Q"
:
=
(
uPred_bupd
Q
)
(
at
level
99
,
Q
at
level
200
,
format
"|==> Q"
)
:
uPred_scope
.
Notation
"P =
=★
Q"
:
=
(
P
⊢
|==>
Q
)
(
at
level
99
,
Q
at
level
200
,
only
parsing
)
:
C_scope
.
Notation
"P =
r
=★ Q"
:
=
(
P
-
★
|=
r
=>
Q
)%
I
(
at
level
99
,
Q
at
level
200
,
format
"P =
r
=★ Q"
)
:
uPred_scope
.
Notation
"P ==★ Q"
:
=
(
P
-
★
|==>
Q
)%
I
(
at
level
99
,
Q
at
level
200
,
format
"P ==★ Q"
)
:
uPred_scope
.
Definition
uPred_iff
{
M
}
(
P
Q
:
uPred
M
)
:
uPred
M
:
=
((
P
→
Q
)
∧
(
Q
→
P
))%
I
.
Instance
:
Params
(@
uPred_iff
)
1
.
...
...
@@ -1283,20 +1283,20 @@ Lemma always_cmra_valid {A : cmraT} (a : A) : □ ✓ a ⊣⊢ ✓ a.
Qed
.
(* Basic update modality *)
Lemma
bupd_intro
P
:
P
=
r
=>
P
.
Lemma
bupd_intro
P
:
P
=
=
★
P
.
Proof
.
unseal
.
split
=>
n
x
?
HP
k
yf
?
;
exists
x
;
split
;
first
done
.
apply
uPred_closed
with
n
;
eauto
using
cmra_validN_op_l
.
Qed
.
Lemma
bupd_mono
P
Q
:
(
P
⊢
Q
)
→
(|=
r
=>
P
)
=
r
=>
Q
.
Lemma
bupd_mono
P
Q
:
(
P
⊢
Q
)
→
(|==>
P
)
=
=
★
Q
.
Proof
.
unseal
.
intros
HPQ
;
split
=>
n
x
?
HP
k
yf
??.
destruct
(
HP
k
yf
)
as
(
x'
&?&?)
;
eauto
.
exists
x'
;
split
;
eauto
using
uPred_in_entails
,
cmra_validN_op_l
.
Qed
.
Lemma
bupd_trans
P
:
(|=
r
=>
|=
r
=>
P
)
=
r
=>
P
.
Lemma
bupd_trans
P
:
(|==>
|==>
P
)
=
=
★
P
.
Proof
.
unseal
;
split
;
naive_solver
.
Qed
.
Lemma
bupd_frame_r
P
R
:
(|=
r
=>
P
)
★
R
=
r
=>
P
★
R
.
Lemma
bupd_frame_r
P
R
:
(|==>
P
)
★
R
=
=
★
P
★
R
.
Proof
.
unseal
;
split
;
intros
n
x
?
(
x1
&
x2
&
Hx
&
HP
&?)
k
yf
??.
destruct
(
HP
k
(
x2
⋅
yf
))
as
(
x'
&?&?)
;
eauto
.
...
...
@@ -1306,7 +1306,7 @@ Proof.
apply
uPred_closed
with
n
;
eauto
3
using
cmra_validN_op_l
,
cmra_validN_op_r
.
Qed
.
Lemma
bupd_ownM_updateP
x
(
Φ
:
M
→
Prop
)
:
x
~~>
:
Φ
→
uPred_ownM
x
=
r
=>
∃
y
,
■
Φ
y
∧
uPred_ownM
y
.
x
~~>
:
Φ
→
uPred_ownM
x
=
=
★
∃
y
,
■
Φ
y
∧
uPred_ownM
y
.
Proof
.
unseal
=>
Hup
;
split
=>
n
x2
?
[
x3
Hx
]
k
yf
??.
destruct
(
Hup
k
(
Some
(
x3
⋅
yf
)))
as
(
y
&?&?)
;
simpl
in
*.
...
...
@@ -1320,20 +1320,20 @@ Global Instance bupd_mono' : Proper ((⊢) ==> (⊢)) (@uPred_bupd M).
Proof
.
intros
P
Q
;
apply
bupd_mono
.
Qed
.
Global
Instance
bupd_flip_mono'
:
Proper
(
flip
(
⊢
)
==>
flip
(
⊢
))
(@
uPred_bupd
M
).
Proof
.
intros
P
Q
;
apply
bupd_mono
.
Qed
.
Lemma
bupd_frame_l
R
Q
:
(
R
★
|=
r
=>
Q
)
=
r
=>
R
★
Q
.
Lemma
bupd_frame_l
R
Q
:
(
R
★
|==>
Q
)
=
=
★
R
★
Q
.
Proof
.
rewrite
!(
comm
_
R
)
;
apply
bupd_frame_r
.
Qed
.
Lemma
bupd_wand_l
P
Q
:
(
P
-
★
Q
)
★
(|=
r
=>
P
)
=
r
=>
Q
.
Lemma
bupd_wand_l
P
Q
:
(
P
-
★
Q
)
★
(|==>
P
)
=
=
★
Q
.
Proof
.
by
rewrite
bupd_frame_l
wand_elim_l
.
Qed
.
Lemma
bupd_wand_r
P
Q
:
(|=
r
=>
P
)
★
(
P
-
★
Q
)
=
r
=>
Q
.
Lemma
bupd_wand_r
P
Q
:
(|==>
P
)
★
(
P
-
★
Q
)
=
=
★
Q
.
Proof
.
by
rewrite
bupd_frame_r
wand_elim_r
.
Qed
.
Lemma
bupd_sep
P
Q
:
(|=
r
=>
P
)
★
(|=
r
=>
Q
)
=
r
=>
P
★
Q
.
Lemma
bupd_sep
P
Q
:
(|==>
P
)
★
(|==>
Q
)
=
=
★
P
★
Q
.
Proof
.
by
rewrite
bupd_frame_r
bupd_frame_l
bupd_trans
.
Qed
.
Lemma
bupd_ownM_update
x
y
:
x
~~>
y
→
uPred_ownM
x
⊢
|=
r
=>
uPred_ownM
y
.
Lemma
bupd_ownM_update
x
y
:
x
~~>
y
→
uPred_ownM
x
⊢
|==>
uPred_ownM
y
.
Proof
.
intros
;
rewrite
(
bupd_ownM_updateP
_
(
y
=))
;
last
by
apply
cmra_update_updateP
.
by
apply
bupd_mono
,
exist_elim
=>
y'
;
apply
pure_elim_l
=>
->.
Qed
.
Lemma
except_last_bupd
P
:
◇
(|=
r
=>
P
)
⊢
(|=
r
=>
◇
P
).
Lemma
except_last_bupd
P
:
◇
(|==>
P
)
⊢
(|==>
◇
P
).
Proof
.
rewrite
/
uPred_except_last
.
apply
or_elim
;
auto
using
bupd_mono
.
by
rewrite
-
bupd_intro
-
or_intro_l
.
...
...
@@ -1490,9 +1490,9 @@ Lemma always_entails_r P Q `{!PersistentP Q} : (P ⊢ Q) → P ⊢ P ★ Q.
Proof
.
by
rewrite
-(
always_always
Q
)
;
apply
always_entails_r'
.
Qed
.
(** Consistency and adequancy statements *)
Lemma
adequacy
φ
n
:
(
True
⊢
Nat
.
iter
n
(
λ
P
,
|=
r
=>
▷
P
)
(
■
φ
))
→
φ
.
Lemma
adequacy
φ
n
:
(
True
⊢
Nat
.
iter
n
(
λ
P
,
|==>
▷
P
)
(
■
φ
))
→
φ
.
Proof
.
cut
(
∀
x
,
✓
{
n
}
x
→
Nat
.
iter
n
(
λ
P
,
|=
r
=>
▷
P
)%
I
(
■
φ
)%
I
n
x
→
φ
).
cut
(
∀
x
,
✓
{
n
}
x
→
Nat
.
iter
n
(
λ
P
,
|==>
▷
P
)%
I
(
■
φ
)%
I
n
x
→
φ
).
{
intros
help
H
.
eapply
(
help
∅
)
;
eauto
using
ucmra_unit_validN
.
eapply
H
;
try
unseal
;
eauto
using
ucmra_unit_validN
.
}
unseal
.
induction
n
as
[|
n
IH
]=>
x
Hx
Hupd
;
auto
.
...
...
@@ -1500,7 +1500,7 @@ Proof.
eapply
IH
with
x'
;
eauto
using
cmra_validN_S
,
cmra_validN_op_l
.
Qed
.
Corollary
consistency_modal
n
:
¬
(
True
⊢
Nat
.
iter
n
(
λ
P
,
|=
r
=>
▷
P
)
False
).
Corollary
consistency_modal
n
:
¬
(
True
⊢
Nat
.
iter
n
(
λ
P
,
|==>
▷
P
)
False
).
Proof
.
exact
(
adequacy
False
n
).
Qed
.
Corollary
consistency
:
¬
(
True
⊢
False
).
...
...
heap_lang/lib/barrier/proof.v
View file @
568f6b7a
...
...
@@ -162,7 +162,7 @@ Proof.
Qed
.
Lemma
recv_split
E
l
P1
P2
:
nclose
N
⊆
E
→
recv
l
(
P1
★
P2
)
={
E
}=
>
recv
l
P1
★
recv
l
P2
.
nclose
N
⊆
E
→
recv
l
(
P1
★
P2
)
={
E
}=
★
recv
l
P1
★
recv
l
P2
.
Proof
.
rename
P1
into
R1
;
rename
P2
into
R2
.
rewrite
{
1
}/
recv
/
barrier_ctx
.
iIntros
(?).
iDestruct
1
as
(
γ
P
Q
i
)
"(#(%&Hh&Hsts)&Hγ&#HQ&HQR)"
.
...
...
heap_lang/lib/barrier/specification.v
View file @
568f6b7a
...
...
@@ -23,7 +23,7 @@ Proof.
-
iIntros
(
P
)
"#? !# _"
.
iApply
(
newbarrier_spec
_
P
)
;
eauto
.
-
iIntros
(
l
P
)
"!# [Hl HP]"
.
by
iApply
signal_spec
;
iFrame
"Hl HP"
.
-
iIntros
(
l
P
)
"!# Hl"
.
iApply
wait_spec
;
iFrame
"Hl"
;
eauto
.
-
intros
;
by
a
pply
recv_split
.
-
i
I
ntros
(
l
P
Q
)
"!#"
.
by
iA
pply
recv_split
.
-
apply
recv_weaken
.
Qed
.
End
spec
.
program_logic/adequacy.v
View file @
568f6b7a
...
...
@@ -40,7 +40,7 @@ Notation wptp t := ([★ list] ef ∈ t, WP ef {{ _, True }})%I.
Lemma
wp_step
e1
σ
1 e2
σ
2
efs
Φ
:
prim_step
e1
σ
1 e2
σ
2
efs
→
world
σ
1
★
WP
e1
{{
Φ
}}
=
r
=>
▷
|=
r
=>
◇
(
world
σ
2
★
WP
e2
{{
Φ
}}
★
wptp
efs
).
world
σ
1
★
WP
e1
{{
Φ
}}
=
=
★
▷
|==>
◇
(
world
σ
2
★
WP
e2
{{
Φ
}}
★
wptp
efs
).
Proof
.
rewrite
{
1
}
wp_unfold
/
wp_pre
.
iIntros
(
Hstep
)
"[(Hw & HE & Hσ) [H|[_ H]]]"
.
{
iDestruct
"H"
as
(
v
)
"[% _]"
.
apply
val_stuck
in
Hstep
;
simplify_eq
.
}
...
...
@@ -54,7 +54,7 @@ Qed.
Lemma
wptp_step
e1
t1
t2
σ
1
σ
2
Φ
:
step
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
world
σ
1
★
WP
e1
{{
Φ
}}
★
wptp
t1
=
r
=>
∃
e2
t2'
,
t2
=
e2
::
t2'
★
▷
|=
r
=>
◇
(
world
σ
2
★
WP
e2
{{
Φ
}}
★
wptp
t2'
).
=
=
★
∃
e2
t2'
,
t2
=
e2
::
t2'
★
▷
|==>
◇
(
world
σ
2
★
WP
e2
{{
Φ
}}
★
wptp
t2'
).
Proof
.
iIntros
(
Hstep
)
"(HW & He & Ht)"
.
destruct
Hstep
as
[
e1'
σ
1
'
e2'
σ
2
'
efs
[|?
t1'
]
t2'
??
Hstep
]
;
simplify_eq
/=.
...
...
@@ -69,7 +69,7 @@ Qed.
Lemma
wptp_steps
n
e1
t1
t2
σ
1
σ
2
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
world
σ
1
★
WP
e1
{{
Φ
}}
★
wptp
t1
⊢
Nat
.
iter
(
S
n
)
(
λ
P
,
|=
r
=>
▷
P
)
(
∃
e2
t2'
,
Nat
.
iter
(
S
n
)
(
λ
P
,
|==>
▷
P
)
(
∃
e2
t2'
,
t2
=
e2
::
t2'
★
world
σ
2
★
WP
e2
{{
Φ
}}
★
wptp
t2'
).
Proof
.
revert
e1
t1
t2
σ
1
σ
2
;
simpl
;
induction
n
as
[|
n
IH
]=>
e1
t1
t2
σ
1
σ
2
/=.
...
...
@@ -79,11 +79,11 @@ Proof.
iUpdIntro
;
iNext
;
iUpd
"H"
as
">?"
.
by
iApply
IH
.
Qed
.
Instance
bupd_iter_mono
n
:
Proper
((
⊢
)
==>
(
⊢
))
(
Nat
.
iter
n
(
λ
P
,
|=
r
=>
▷
P
)%
I
).
Instance
bupd_iter_mono
n
:
Proper
((
⊢
)
==>
(
⊢
))
(
Nat
.
iter
n
(
λ
P
,
|==>
▷
P
)%
I
).
Proof
.
intros
P
Q
HP
.
induction
n
;
simpl
;
do
2
?f_equiv
;
auto
.
Qed
.
Lemma
bupd_iter_frame_l
n
R
Q
:
R
★
Nat
.
iter
n
(
λ
P
,
|=
r
=>
▷
P
)
Q
⊢
Nat
.
iter
n
(
λ
P
,
|=
r
=>
▷
P
)
(
R
★
Q
).
R
★
Nat
.
iter
n
(
λ
P
,
|==>
▷
P
)
Q
⊢
Nat
.
iter
n
(
λ
P
,
|==>
▷
P
)
(
R
★
Q
).
Proof
.
induction
n
as
[|
n
IH
]
;
simpl
;
[
done
|].
by
rewrite
bupd_frame_l
{
1
}(
later_intro
R
)
-
later_sep
IH
.
...
...
@@ -92,7 +92,7 @@ Qed.
Lemma
wptp_result
n
e1
t1
v2
t2
σ
1
σ
2
φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
of_val
v2
::
t2
,
σ
2
)
→
world
σ
1
★
WP
e1
{{
v
,
■
φ
v
}}
★
wptp
t1
⊢
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|=
r
=>
▷
P
)
(
■
φ
v2
).
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|==>
▷
P
)
(
■
φ
v2
).
Proof
.
intros
.
rewrite
wptp_steps
//.
rewrite
(
Nat_iter_S_r
(
S
n
)).
apply
bupd_iter_mono
.
...
...
@@ -102,7 +102,7 @@ Proof.
Qed
.
Lemma
wp_safe
e
σ
Φ
:
world
σ
★
WP
e
{{
Φ
}}
=
r
=>
▷
■
(
is_Some
(
to_val
e
)
∨
reducible
e
σ
).
world
σ
★
WP
e
{{
Φ
}}
=
=
★
▷
■
(
is_Some
(
to_val
e
)
∨
reducible
e
σ
).
Proof
.
rewrite
wp_unfold
/
wp_pre
.
iIntros
"[(Hw&HE&Hσ) [H|[_ H]]]"
.
{
iDestruct
"H"
as
(
v
)
"[% _]"
;
eauto
10
.
}
...
...
@@ -113,7 +113,7 @@ Qed.
Lemma
wptp_safe
n
e1
e2
t1
t2
σ
1
σ
2
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
e2
∈
t2
→
world
σ
1
★
WP
e1
{{
Φ
}}
★
wptp
t1
⊢
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|=
r
=>
▷
P
)
(
■
(
is_Some
(
to_val
e2
)
∨
reducible
e2
σ
2
)).
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|==>
▷
P
)
(
■
(
is_Some
(
to_val
e2
)
∨
reducible
e2
σ
2
)).
Proof
.
intros
?
He2
.
rewrite
wptp_steps
//
;
rewrite
(
Nat_iter_S_r
(
S
n
)).
apply
bupd_iter_mono
.
iDestruct
1
as
(
e2'
t2'
)
"(% & Hw & H & Htp)"
;
simplify_eq
.
...
...
@@ -123,9 +123,9 @@ Qed.
Lemma
wptp_invariance
n
e1
e2
t1
t2
σ
1
σ
2
I
φ
Φ
:
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
(
I
={
⊤
,
∅
}=
>
∃
σ
'
,
ownP
σ
'
∧
■
φ
σ
'
)
→
(
I
={
⊤
,
∅
}=
★
∃
σ
'
,
ownP
σ
'
∧
■
φ
σ
'
)
→
I
★
world
σ
1
★
WP
e1
{{
Φ
}}
★
wptp
t1
⊢
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|=
r
=>
▷
P
)
(
■
φ
σ
2
).
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|==>
▷
P
)
(
■
φ
σ
2
).
Proof
.
intros
?
HI
.
rewrite
wptp_steps
//.
rewrite
(
Nat_iter_S_r
(
S
n
))
bupd_iter_frame_l
.
apply
bupd_iter_mono
.
...
...
@@ -156,8 +156,8 @@ Proof.
Qed
.
Theorem
wp_invariance
Σ
`
{
irisPreG
Λ
Σ
}
e
σ
1
t2
σ
2
I
φ
Φ
:
(
∀
`
{
irisG
Λ
Σ
},
ownP
σ
1
={
⊤
}=
>
I
★
WP
e
{{
Φ
}})
→
(
∀
`
{
irisG
Λ
Σ
},
I
={
⊤
,
∅
}=
>
∃
σ
'
,
ownP
σ
'
∧
■
φ
σ
'
)
→
(
∀
`
{
irisG
Λ
Σ
},
ownP
σ
1
={
⊤
}=
★
I
★
WP
e
{{
Φ
}})
→
(
∀
`
{
irisG
Λ
Σ
},
I
={
⊤
,
∅
}=
★
∃
σ
'
,
ownP
σ
'
∧
■
φ
σ
'
)
→
rtc
step
([
e
],
σ
1
)
(
t2
,
σ
2
)
→
φ
σ
2
.
Proof
.
...
...
program_logic/auth.v
View file @
568f6b7a
...
...
@@ -91,7 +91,7 @@ Section auth.
Proof
.
intros
a1
a2
.
apply
auth_own_mono
.
Qed
.
Lemma
auth_alloc_strong
N
E
t
(
G
:
gset
gname
)
:
✓
(
f
t
)
→
▷
φ
t
={
E
}=
>
∃
γ
,
■
(
γ
∉
G
)
∧
auth_ctx
γ
N
f
φ
∧
auth_own
γ
(
f
t
).
✓
(
f
t
)
→
▷
φ
t
={
E
}=
★
∃
γ
,
■
(
γ
∉
G
)
∧
auth_ctx
γ
N
f
φ
∧
auth_own
γ
(
f
t
).
Proof
.
iIntros
(?)
"Hφ"
.
rewrite
/
auth_own
/
auth_ctx
.
iUpd
(
own_alloc_strong
(
Auth
(
Excl'
(
f
t
))
(
f
t
))
G
)
as
(
γ
)
"[% Hγ]"
;
first
done
.
...
...
@@ -102,17 +102,17 @@ Section auth.
Qed
.
Lemma
auth_alloc
N
E
t
:
✓
(
f
t
)
→
▷
φ
t
={
E
}=
>
∃
γ
,
auth_ctx
γ
N
f
φ
∧
auth_own
γ
(
f
t
).
✓
(
f
t
)
→
▷
φ
t
={
E
}=
★
∃
γ
,
auth_ctx
γ
N
f
φ
∧
auth_own
γ
(
f
t
).
Proof
.
iIntros
(?)
"Hφ"
.
iUpd
(
auth_alloc_strong
N
E
t
∅
with
"Hφ"
)
as
(
γ
)
"[_ ?]"
;
eauto
.
Qed
.
Lemma
auth_empty
γ
:
True
=
r
=>
auth_own
γ
∅
.
Lemma
auth_empty
γ
:
True
=
=
★
auth_own
γ
∅
.
Proof
.
by
rewrite
/
auth_own
-
own_empty
.
Qed
.
Lemma
auth_acc
E
γ
a
:
▷
auth_inv
γ
f
φ
★
auth_own
γ
a
={
E
}=
>
∃
t
,
▷
auth_inv
γ
f
φ
★
auth_own
γ
a
={
E
}=
★
∃
t
,
■
(
a
≼
f
t
)
★
▷
φ
t
★
∀
u
b
,
■
((
f
t
,
a
)
~l
~>
(
f
u
,
b
))
★
▷
φ
u
={
E
}=
★
▷
auth_inv
γ
f
φ
★
auth_own
γ
b
.
Proof
.
...
...
@@ -128,7 +128,7 @@ Section auth.
Lemma
auth_open
E
N
γ
a
:
nclose
N
⊆
E
→
auth_ctx
γ
N
f
φ
★
auth_own
γ
a
={
E
,
E
∖
N
}=
>
∃
t
,
auth_ctx
γ
N
f
φ
★
auth_own
γ
a
={
E
,
E
∖
N
}=
★
∃
t
,
■
(
a
≼
f
t
)
★
▷
φ
t
★
∀
u
b
,
■
((
f
t
,
a
)
~l
~>
(
f
u
,
b
))
★
▷
φ
u
={
E
∖
N
,
E
}=
★
auth_own
γ
b
.
Proof
.
...
...
program_logic/boxes.v
View file @
568f6b7a
...
...
@@ -63,7 +63,7 @@ Qed.
Lemma
box_own_auth_update
γ
b1
b2
b3
:
box_own_auth
γ
(
●
Excl'
b1
)
★
box_own_auth
γ
(
◯
Excl'
b2
)
=
r
=>
box_own_auth
γ
(
●
Excl'
b3
)
★
box_own_auth
γ
(
◯
Excl'
b3
).
=
=
★
box_own_auth
γ
(
●
Excl'
b3
)
★
box_own_auth
γ
(
◯
Excl'
b3
).
Proof
.
rewrite
/
box_own_auth
-!
own_op
.
apply
own_update
,
prod_update
;
last
done
.
by
apply
auth_update
,
option_local_update
,
exclusive_local_update
.
...
...
@@ -86,7 +86,7 @@ Proof.
Qed
.
Lemma
box_insert
f
P
Q
:
▷
box
N
f
P
={
N
}=
>
∃
γ
,
f
!!
γ
=
None
★
▷
box
N
f
P
={
N
}=
★
∃
γ
,
f
!!
γ
=
None
★
slice
N
γ
Q
★
▷
box
N
(<[
γ
:
=
false
]>
f
)
(
Q
★
P
).
Proof
.
iDestruct
1
as
(
Φ
)
"[#HeqP Hf]"
.
...
...
@@ -106,7 +106,7 @@ Qed.
Lemma
box_delete
f
P
Q
γ
:
f
!!
γ
=
Some
false
→
slice
N
γ
Q
★
▷
box
N
f
P
={
N
}=
>
∃
P'
,
slice
N
γ
Q
★
▷
box
N
f
P
={
N
}=
★
∃
P'
,
▷
▷
(
P
≡
(
Q
★
P'
))
★
▷
box
N
(
delete
γ
f
)
P'
.
Proof
.
iIntros
(?)
"[#Hinv H]"
;
iDestruct
"H"
as
(
Φ
)
"[#HeqP Hf]"
.
...
...
@@ -125,7 +125,7 @@ Qed.
Lemma
box_fill
f
γ
P
Q
:
f
!!
γ
=
Some
false
→
slice
N
γ
Q
★
▷
Q
★
▷
box
N
f
P
={
N
}=
>
▷
box
N
(<[
γ
:
=
true
]>
f
)
P
.
slice
N
γ
Q
★
▷
Q
★
▷
box
N
f
P
={
N
}=
★
▷
box
N
(<[
γ
:
=
true
]>
f
)
P
.
Proof
.
iIntros
(?)
"(#Hinv & HQ & H)"
;
iDestruct
"H"
as
(
Φ
)
"[#HeqP Hf]"
.
iInv
N
as
(
b'
)
"(>Hγ & #HγQ & _)"
"Hclose"
.
...
...
@@ -143,7 +143,7 @@ Qed.
Lemma
box_empty
f
P
Q
γ
:
f
!!
γ
=
Some
true
→
slice
N
γ
Q
★
▷
box
N
f
P
={
N
}=
>
▷
Q
★
▷
box
N
(<[
γ
:
=
false
]>
f
)
P
.
slice
N
γ
Q
★
▷
box
N
f
P
={
N
}=
★
▷
Q
★
▷
box
N
(<[
γ
:
=
false
]>
f
)
P
.
Proof
.
iIntros
(?)
"[#Hinv H]"
;
iDestruct
"H"
as
(
Φ
)
"[#HeqP Hf]"
.
iInv
N
as
(
b
)
"(>Hγ & #HγQ & HQ)"
"Hclose"
.
...
...
@@ -160,7 +160,7 @@ Proof.
iFrame
;
eauto
.
Qed
.
Lemma
box_fill_all
f
P
Q
:
box
N
f
P
★
▷
P
={
N
}=
>
box
N
(
const
true
<$>
f
)
P
.
Lemma
box_fill_all
f
P
Q
:
box
N
f
P
★
▷
P
={
N
}=
★
box
N
(
const
true
<$>
f
)
P
.
Proof
.
iIntros
"[H HP]"
;
iDestruct
"H"
as
(
Φ
)
"[#HeqP Hf]"
.
iExists
Φ
;
iSplitR
;
first
by
rewrite
big_sepM_fmap
.
...
...
@@ -176,7 +176,7 @@ Qed.
Lemma
box_empty_all
f
P
Q
:
map_Forall
(
λ
_
,
(
true
=))
f
→
box
N
f
P
={
N
}=
>
▷
P
★
box
N
(
const
false
<$>
f
)
P
.
box
N
f
P
={
N
}=
★
▷
P
★
box
N
(
const
false
<$>
f
)
P
.
Proof
.
iDestruct
1
as
(
Φ
)
"[#HeqP Hf]"
.
iAssert
([
★
map
]
γ↦
b
∈
f
,
▷
Φ
γ
★
box_own_auth
γ
(
◯
Excl'
false
)
★
...
...
program_logic/cancelable_invariants.v
View file @
568f6b7a
...
...
@@ -44,7 +44,7 @@ Section proofs.
Lemma
cinv_own_1_l
γ
q
:
cinv_own
γ
1
★
cinv_own
γ
q
⊢
False
.
Proof
.
rewrite
cinv_own_valid
.
by
iIntros
(?%(
exclusive_l
1
%
Qp
)).
Qed
.
Lemma
cinv_alloc
E
N
P
:
▷
P
={
E
}=
>
∃
γ
,
cinv
N
γ
P
★
cinv_own
γ
1
.
Lemma
cinv_alloc
E
N
P
:
▷
P
={
E
}=
★
∃
γ
,
cinv
N
γ
P
★
cinv_own
γ
1
.
Proof
.
rewrite
/
cinv
/
cinv_own
.
iIntros
"HP"
.
iUpd
(
own_alloc
1
%
Qp
)
as
(
γ
)
"H1"
;
first
done
.
...
...
program_logic/counter_examples.v
View file @
568f6b7a
...
...
@@ -13,13 +13,13 @@ Module savedprop. Section savedprop.
Context
(
sprop
:
Type
)
(
saved
:
sprop
→
iProp
→
iProp
).
Hypothesis
sprop_persistent
:
∀
i
P
,
PersistentP
(
saved
i
P
).
Hypothesis
sprop_alloc_dep
:
∀
(
P
:
sprop
→
iProp
),
True
=
r
=>
(
∃
i
,
saved
i
(
P
i
)).
∀
(
P
:
sprop
→
iProp
),
True
=
=
★
(
∃
i
,
saved
i
(
P
i
)).
Hypothesis
sprop_agree
:
∀
i
P
Q
,
saved
i
P
∧
saved
i
Q
⊢
□
(
P
↔
Q
).
(** A bad recursive reference: "Assertion with name [i] does not hold" *)
Definition
A
(
i
:
sprop
)
:
iProp
:
=
∃
P
,
¬
P
★
saved
i
P
.
Lemma
A_alloc
:
True
=
r
=>
∃
i
,
saved
i
(
A
i
).
Lemma
A_alloc
:
True
=
=
★
∃
i
,
saved
i
(
A
i
).
Proof
.
by
apply
sprop_alloc_dep
.
Qed
.
Lemma
saved_NA
i
:
saved
i
(
A
i
)
⊢
¬
A
i
.
...
...
program_logic/fancy_updates.v
View file @
568f6b7a
...
...
@@ -6,7 +6,7 @@ Import uPred.
Program
Definition
fupd_def
`
{
irisG
Λ
Σ
}
(
E1
E2
:
coPset
)
(
P
:
iProp
Σ
)
:
iProp
Σ
:
=
(
wsat
★
ownE
E1
=
r
=
★
◇
(
wsat
★
ownE
E2
★
P
))%
I
.
(
wsat
★
ownE
E1
==
★
◇
(
wsat
★
ownE
E2
★
P
))%
I
.
Definition
fupd_aux
:
{
x
|
x
=
@
fupd_def
}.
by
eexists
.
Qed
.
Definition
fupd
:
=
proj1_sig
fupd_aux
.
Definition
fupd_eq
:
@
fupd
=
@
fupd_def
:
=
proj2_sig
fupd_aux
.
...
...
@@ -19,7 +19,7 @@ Notation "|={ E1 , E2 }=> Q" := (fupd E1 E2 Q)
Notation
"P ={ E1 , E2 }=★ Q"
:
=
(
P
-
★
|={
E1
,
E2
}=>
Q
)%
I
(
at
level
99
,
E1
,
E2
at
level
50
,
Q
at
level
200
,
format
"P ={ E1 , E2 }=★ Q"
)
:
uPred_scope
.
Notation
"P ={ E1 , E2 }=
>
Q"
:
=
(
P
⊢
|={
E1
,
E2
}=>
Q
)
Notation
"P ={ E1 , E2 }=
★
Q"
:
=
(
P
⊢
|={
E1
,
E2
}=>
Q
)
(
at
level
99
,
E1
,
E2
at
level
50
,
Q
at
level
200
,
only
parsing
)
:
C_scope
.
Notation
"|={ E }=> Q"
:
=
(
fupd
E
E
Q
)
...
...
@@ -28,7 +28,7 @@ Notation "|={ E }=> Q" := (fupd E E Q)
Notation
"P ={ E }=★ Q"
:
=
(
P
-
★
|={
E
}=>
Q
)%
I
(
at
level
99
,
E
at
level
50
,
Q
at
level
200
,
format
"P ={ E }=★ Q"
)
:
uPred_scope
.
Notation
"P ={ E }=
>
Q"
:
=
(
P
⊢
|={
E
}=>
Q
)
Notation
"P ={ E }=
★
Q"
:
=
(
P
⊢
|={
E
}=>
Q
)
(
at
level
99
,
E
at
level
50
,
Q
at
level
200
,
only
parsing
)
:
C_scope
.
Notation
"|={ E1 , E2 }▷=> Q"
:
=
(|={
E1
,
E2
}=>
(
▷
|={
E2
,
E1
}=>
Q
))%
I
...
...
@@ -55,31 +55,31 @@ Proof.
by
iFrame
.
Qed
.
Lemma
except_last_fupd
E1
E2
P
:
◇
(|={
E1
,
E2
}=>
P
)
={
E1
,
E2
}=
>
P
.
Lemma
except_last_fupd
E1
E2
P
:
◇
(|={
E1
,
E2
}=>
P
)
={
E1
,
E2
}=
★
P
.
Proof
.
rewrite
fupd_eq
.
iIntros
"H [Hw HE]"
.
iTimeless
"H"
.
iApply
"H"
;
by
iFrame
.
Qed
.
Lemma
bupd_fupd
E
P
:
(|=
r
=>
P
)
={
E
}=
>
P
.
Lemma
bupd_fupd
E
P
:
(|==>
P
)
={
E
}=
★
P
.
Proof
.
rewrite
fupd_eq
/
fupd_def
.
iIntros
"H [$ $]"
;
iUpd
"H"
.
iUpdIntro
.
by
iApply
except_last_intro
.
Qed
.
Lemma
fupd_mono
E1
E2
P
Q
:
(
P
⊢
Q
)
→
(|={
E1
,
E2
}=>
P
)
={
E1
,
E2
}=
>
Q
.
Lemma
fupd_mono
E1
E2
P
Q
:
(
P
⊢
Q
)
→
(|={
E1
,
E2
}=>
P
)
={
E1
,
E2
}=
★
Q
.
Proof
.
rewrite
fupd_eq
/
fupd_def
.
iIntros
(
HPQ
)
"HP HwE"
.
rewrite
-
HPQ
.
by
iApply
"HP"
.
Qed
.
Lemma
fupd_trans
E1
E2
E3
P
:
(|={
E1
,
E2
}=>
|={
E2
,
E3
}=>
P
)
={
E1
,
E3
}=
>
P
.
Lemma
fupd_trans
E1
E2
E3
P
:
(|={
E1
,
E2
}=>
|={
E2
,
E3
}=>
P
)
={
E1
,
E3
}=
★
P
.
Proof
.
rewrite
fupd_eq
/
fupd_def
.
iIntros
"HP HwE"
.
iUpd
(
"HP"
with
"HwE"
)
as
">(Hw & HE & HP)"
.
iApply
"HP"
;
by
iFrame
.
Qed
.
Lemma
fupd_mask_frame_r'
E1
E2
Ef
P
:
E1
⊥
Ef
→
(|={
E1
,
E2
}=>
E2
⊥
Ef
→
P
)
={
E1
∪
Ef
,
E2
∪
Ef
}=
>
P
.
E1
⊥
Ef
→
(|={
E1
,
E2
}=>
E2
⊥
Ef
→
P
)
={
E1
∪
Ef
,
E2
∪
Ef
}=
★
P
.
Proof
.
intros
.
rewrite
fupd_eq
/
fupd_def
ownE_op
//.
iIntros
"Hvs (Hw & HE1 &HEf)"
.
iUpd
(
"Hvs"
with
"[Hw HE1]"
)
as
">($ & HE2 & HP)"
;
first
by
iFrame
.
...
...
@@ -87,7 +87,7 @@ Proof.
iUpdIntro
;
iApply
except_last_intro
.
by
iApply
"HP"
.
Qed
.
Lemma
fupd_frame_r
E1
E2
P
Q
:
(|={
E1
,
E2
}=>
P
)
★
Q
={
E1
,
E2
}=
>
P
★
Q
.
Lemma
fupd_frame_r
E1
E2
P
Q
:
(|={
E1
,
E2
}=>
P
)
★
Q
={
E1
,
E2
}=
★
P
★
Q
.
Proof
.
rewrite
fupd_eq
/
fupd_def
.
by
iIntros
"[HwP $]"
.
Qed
.
(** * Derived rules *)
...
...
@@ -97,50 +97,50 @@ Global Instance fupd_flip_mono' E1 E2 :
Proper
(
flip
(
⊢
)
==>
flip
(
⊢
))
(@
fupd
Λ
Σ
_
E1
E2
).
Proof
.
intros
P
Q
;
apply
fupd_mono
.
Qed
.
Lemma
fupd_intro
E
P
:
P
={
E
}=
>
P
.
Lemma
fupd_intro
E
P
:
P
={
E
}=
★
P
.
Proof
.
iIntros
"HP"
.
by
iApply
bupd_fupd
.
Qed
.
Lemma
fupd_intro_mask'
E1
E2
:
E2
⊆
E1
→
True
⊢
|={
E1
,
E2
}=>
|={
E2
,
E1
}=>
True
.
Proof
.
exact
:
fupd_intro_mask
.
Qed
.
Lemma
fupd_except_last
E1
E2
P
:
(|={
E1
,
E2
}=>
◇
P
)
={
E1
,
E2
}=
>
P
.
Lemma
fupd_except_last
E1
E2
P
:
(|={
E1
,
E2
}=>
◇
P
)
={
E1
,
E2
}=
★
P
.
Proof
.
by
rewrite
{
1
}(
fupd_intro
E2
P
)
except_last_fupd
fupd_trans
.
Qed
.
Lemma
fupd_frame_l
E1
E2
P
Q
:
(
P
★
|={
E1
,
E2
}=>
Q
)
={
E1
,
E2
}=
>
P
★
Q
.
Lemma
fupd_frame_l
E1
E2
P
Q
:
(
P
★
|={
E1
,
E2
}=>
Q
)
={
E1
,
E2
}=
★
P
★
Q
.
Proof
.
rewrite
!(
comm
_
P
)
;
apply
fupd_frame_r
.
Qed
.
Lemma
fupd_wand_l
E1
E2
P
Q
:
(
P
-
★
Q
)
★
(|={
E1
,
E2
}=>
P
)
={
E1
,
E2
}=
>
Q
.
Lemma
fupd_wand_l
E1
E2
P
Q
:
(
P
-
★
Q
)
★
(|={
E1
,
E2
}=>
P
)
={
E1
,
E2
}=
★
Q
.
Proof
.
by
rewrite
fupd_frame_l
wand_elim_l
.
Qed
.
Lemma
fupd_wand_r
E1
E2
P
Q
:
(|={
E1
,
E2
}=>
P
)
★
(
P
-
★
Q
)
={
E1
,
E2
}=
>
Q
.
Lemma
fupd_wand_r
E1
E2
P
Q
:
(|={
E1
,
E2
}=>
P
)
★
(
P
-
★
Q
)
={
E1
,
E2
}=
★
Q
.
Proof
.
by
rewrite
fupd_frame_r
wand_elim_r
.
Qed
.
Lemma
fupd_trans_frame
E1
E2
E3
P
Q
:
((
Q
={
E2
,
E3
}=
★
True
)
★
|={
E1
,
E2
}=>
(
Q
★
P
))
={
E1
,
E3
}=
>
P
.
((
Q
={
E2
,
E3
}=
★
True
)
★
|={
E1
,
E2
}=>
(
Q
★
P
))
={
E1
,
E3
}=
★
P
.
Proof
.
rewrite
fupd_frame_l
assoc
-(
comm
_
Q
)
wand_elim_r
.
by
rewrite
fupd_frame_r
left_id
fupd_trans
.
Qed
.
Lemma
fupd_mask_frame_r
E1
E2
Ef
P
:
E1
⊥
Ef
→
(|={
E1
,
E2
}=>
P
)
={
E1
∪
Ef
,
E2
∪
Ef
}=
>
P
.
E1
⊥
Ef
→
(|={
E1
,
E2
}=>
P
)
={
E1
∪
Ef
,
E2
∪
Ef
}=
★
P
.
Proof
.
iIntros
(?)
"H"
.
iApply
fupd_mask_frame_r'
;
auto
.
iApply
fupd_wand_r
;
iFrame
"H"
;
eauto
.
Qed
.
Lemma
fupd_mask_mono
E1
E2
P
:
E1
⊆
E2
→
(|={
E1
}=>
P
)
={
E2
}=
>
P
.
Lemma
fupd_mask_mono
E1
E2
P
:
E1
⊆
E2
→
(|={
E1
}=>
P
)
={
E2
}=
★
P
.
Proof
.
intros
(
Ef
&->&?)%
subseteq_disjoint_union_L
.
by
apply
fupd_mask_frame_r
.
Qed
.
Lemma
fupd_sep
E
P
Q
:
(|={
E
}=>
P
)
★
(|={
E
}=>
Q
)
={
E
}=
>
P
★
Q
.
Lemma
fupd_sep
E
P
Q
:
(|={
E
}=>
P
)
★
(|={
E
}=>
Q
)
={
E
}=
★
P
★
Q
.
Proof
.
by
rewrite
fupd_frame_r
fupd_frame_l
fupd_trans
.
Qed
.
Lemma
fupd_big_sepM
`
{
Countable
K
}
{
A
}
E
(
Φ
:
K
→
A
→
iProp
Σ
)
(
m
:
gmap
K
A
)
:
([
★
map
]
k
↦
x
∈
m
,
|={
E
}=>
Φ
k
x
)
={
E
}=
>
[
★
map
]
k
↦
x
∈
m
,
Φ
k
x
.
([
★
map
]
k
↦
x
∈
m
,
|={
E
}=>
Φ
k
x
)
={
E
}=
★
[
★
map
]
k
↦
x
∈
m
,
Φ
k
x
.
Proof
.
apply
(
big_opM_forall
(
λ
P
Q
,
P
={
E
}=
>
Q
))
;
auto
using
fupd_intro
.
apply
(
big_opM_forall
(
λ
P
Q
,
P
={
E
}=
★
Q
))
;
auto
using
fupd_intro
.
intros
P1
P2
HP
Q1
Q2
HQ
.
by
rewrite
HP
HQ
-
fupd_sep
.
Qed
.
Lemma
fupd_big_sepS
`
{
Countable
A
}
E
(
Φ
:
A
→
iProp
Σ
)
X
:
([
★
set
]
x
∈
X
,
|={
E
}=>
Φ
x
)
={
E
}=
>
[
★
set
]
x
∈
X
,
Φ
x
.
([
★
set
]
x
∈
X
,
|={
E
}=>
Φ
x
)
={
E
}=
★
[
★
set
]
x
∈
X
,
Φ
x
.
Proof
.
apply
(
big_opS_forall
(
λ
P
Q
,
P
={
E
}=
>
Q
))
;
auto
using
fupd_intro
.
apply
(
big_opS_forall
(
λ
P
Q
,
P
={
E
}=
★
Q
))
;
auto
using
fupd_intro
.
intros
P1
P2
HP
Q1
Q2
HQ
.
by
rewrite
HP
HQ
-
fupd_sep
.
Qed
.
End
fupd
.
...
...
@@ -154,7 +154,7 @@ Section proofmode_classes.
Proof
.
rewrite
/
FromPure
.
intros
<-.
apply
fupd_intro
.
Qed
.
Global
Instance
<