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
Iris
Iris
Commits
683b7066
Commit
683b7066
authored
Dec 06, 2016
by
Ralf Jung
Browse files
rename thread-local invariants -> non-atomic invariants
parent
124a7d8d
Pipeline
#3224
passed with stage
in 10 minutes and 55 seconds
Changes
2
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
_CoqProject
View file @
683b7066
...
...
@@ -80,7 +80,7 @@ base_logic/lib/viewshifts.v
base_logic/lib/auth.v
base_logic/lib/sts.v
base_logic/lib/boxes.v
base_logic/lib/
thread_local
.v
base_logic/lib/
na_invariants
.v
base_logic/lib/cancelable_invariants.v
base_logic/lib/counter_examples.v
base_logic/lib/fractional.v
...
...
base_logic/lib/
thread_local
.v
→
base_logic/lib/
na_invariants
.v
View file @
683b7066
...
...
@@ -3,55 +3,57 @@ From iris.algebra Require Export gmap gset coPset.
From
iris
.
proofmode
Require
Import
tactics
.
Import
uPred
.
(* Non-atomic ("thread-local") invariants. *)
Definition
thread_id
:
=
gname
.
Class
thread_local
G
Σ
:
=
Class
na_inv
G
Σ
:
=
tl_inG
:
>
inG
Σ
(
prodR
coPset_disjR
(
gset_disjR
positive
)).
Section
defs
.
Context
`
{
invG
Σ
,
thread_local
G
Σ
}.
Context
`
{
invG
Σ
,
na_inv
G
Σ
}.
Definition
tl
_own
(
tid
:
thread_id
)
(
E
:
coPset
)
:
iProp
Σ
:
=
Definition
na
_own
(
tid
:
thread_id
)
(
E
:
coPset
)
:
iProp
Σ
:
=
own
tid
(
CoPset
E
,
∅
).
Definition
tl
_inv
(
tid
:
thread_id
)
(
N
:
namespace
)
(
P
:
iProp
Σ
)
:
iProp
Σ
:
=
Definition
na
_inv
(
tid
:
thread_id
)
(
N
:
namespace
)
(
P
:
iProp
Σ
)
:
iProp
Σ
:
=
(
∃
i
,
⌜
i
∈
↑
N
⌝
∧
inv
N
(
P
∗
own
tid
(
∅
,
GSet
{[
i
]})
∨
tl
_own
tid
{[
i
]}))%
I
.
inv
N
(
P
∗
own
tid
(
∅
,
GSet
{[
i
]})
∨
na
_own
tid
{[
i
]}))%
I
.
End
defs
.
Instance
:
Params
(@
tl
_inv
)
3
.
Typeclasses
Opaque
tl
_own
tl
_inv
.
Instance
:
Params
(@
na
_inv
)
3
.
Typeclasses
Opaque
na
_own
na
_inv
.
Section
proofs
.
Context
`
{
invG
Σ
,
thread_local
G
Σ
}.
Context
`
{
invG
Σ
,
na_inv
G
Σ
}.
Global
Instance
tl
_own_timeless
tid
E
:
TimelessP
(
tl
_own
tid
E
).
Proof
.
rewrite
/
tl
_own
;
apply
_
.
Qed
.
Global
Instance
na
_own_timeless
tid
E
:
TimelessP
(
na
_own
tid
E
).
Proof
.
rewrite
/
na
_own
;
apply
_
.
Qed
.
Global
Instance
tl
_inv_ne
tid
N
n
:
Proper
(
dist
n
==>
dist
n
)
(
tl
_inv
tid
N
).
Proof
.
rewrite
/
tl
_inv
.
solve_proper
.
Qed
.
Global
Instance
tl
_inv_proper
tid
N
:
Proper
((
≡
)
==>
(
≡
))
(
tl
_inv
tid
N
).
Global
Instance
na
_inv_ne
tid
N
n
:
Proper
(
dist
n
==>
dist
n
)
(
na
_inv
tid
N
).
Proof
.
rewrite
/
na
_inv
.
solve_proper
.
Qed
.
Global
Instance
na
_inv_proper
tid
N
:
Proper
((
≡
)
==>
(
≡
))
(
na
_inv
tid
N
).
Proof
.
apply
(
ne_proper
_
).
Qed
.
Global
Instance
tl
_inv_persistent
tid
N
P
:
PersistentP
(
tl
_inv
tid
N
P
).
Proof
.
rewrite
/
tl
_inv
;
apply
_
.
Qed
.
Global
Instance
na
_inv_persistent
tid
N
P
:
PersistentP
(
na
_inv
tid
N
P
).
Proof
.
rewrite
/
na
_inv
;
apply
_
.
Qed
.
Lemma
tl
_alloc
:
(|==>
∃
tid
,
tl
_own
tid
⊤
)%
I
.
Lemma
na
_alloc
:
(|==>
∃
tid
,
na
_own
tid
⊤
)%
I
.
Proof
.
by
apply
own_alloc
.
Qed
.
Lemma
tl
_own_disjoint
tid
E1
E2
:
tl
_own
tid
E1
-
∗
tl
_own
tid
E2
-
∗
⌜
E1
⊥
E2
⌝
.
Lemma
na
_own_disjoint
tid
E1
E2
:
na
_own
tid
E1
-
∗
na
_own
tid
E2
-
∗
⌜
E1
⊥
E2
⌝
.
Proof
.
apply
wand_intro_r
.
rewrite
/
tl
_own
-
own_op
own_valid
-
coPset_disj_valid_op
.
by
iIntros
([?
_
]).
rewrite
/
na
_own
-
own_op
own_valid
-
coPset_disj_valid_op
.
by
iIntros
([?
_
]).
Qed
.
Lemma
tl
_own_union
tid
E1
E2
:
E1
⊥
E2
→
tl
_own
tid
(
E1
∪
E2
)
⊣
⊢
tl
_own
tid
E1
∗
tl
_own
tid
E2
.
Lemma
na
_own_union
tid
E1
E2
:
E1
⊥
E2
→
na
_own
tid
(
E1
∪
E2
)
⊣
⊢
na
_own
tid
E1
∗
na
_own
tid
E2
.
Proof
.
intros
?.
by
rewrite
/
tl
_own
-
own_op
pair_op
left_id
coPset_disj_union
.
intros
?.
by
rewrite
/
na
_own
-
own_op
pair_op
left_id
coPset_disj_union
.
Qed
.
Lemma
tl
_inv_alloc
tid
E
N
P
:
▷
P
={
E
}=
∗
tl
_inv
tid
N
P
.
Lemma
na
_inv_alloc
tid
E
N
P
:
▷
P
={
E
}=
∗
na
_inv
tid
N
P
.
Proof
.
iIntros
"HP"
.
iMod
(
own_empty
(
prodUR
coPset_disjUR
(
gset_disjUR
positive
))
tid
)
as
"Hempty"
.
...
...
@@ -64,20 +66,20 @@ Section proofs.
eapply
nclose_infinite
,
(
difference_finite_inv
_
_
),
Hfin
.
apply
of_gset_finite
.
}
simpl
.
iDestruct
"Hm"
as
%(<-
&
i
&
->
&
?).
rewrite
/
tl
_inv
.
rewrite
/
na
_inv
.
iMod
(
inv_alloc
N
with
"[-]"
)
;
last
(
iModIntro
;
iExists
i
;
eauto
).
iNext
.
iLeft
.
by
iFrame
.
Qed
.
Lemma
tl
_inv_open
tid
E
N
P
:
Lemma
na
_inv_open
tid
E
N
P
:
↑
N
⊆
E
→
tl
_inv
tid
N
P
-
∗
tl
_own
tid
E
={
E
}=
∗
▷
P
∗
tl
_own
tid
(
E
∖↑
N
)
∗
(
▷
P
∗
tl
_own
tid
(
E
∖↑
N
)
={
E
}=
∗
tl
_own
tid
E
).
na
_inv
tid
N
P
-
∗
na
_own
tid
E
={
E
}=
∗
▷
P
∗
na
_own
tid
(
E
∖↑
N
)
∗
(
▷
P
∗
na
_own
tid
(
E
∖↑
N
)
={
E
}=
∗
na
_own
tid
E
).
Proof
.
rewrite
/
tl
_inv
.
iIntros
(?)
"#Htlinv Htoks"
.
rewrite
/
na
_inv
.
iIntros
(?)
"#Htlinv Htoks"
.
iDestruct
"Htlinv"
as
(
i
)
"[% Hinv]"
.
rewrite
[
E
as
X
in
tl
_own
tid
X
](
union_difference_L
(
↑
N
)
E
)
//.
rewrite
[
X
in
(
X
∪
_
)](
union_difference_L
{[
i
]}
(
↑
N
))
?
tl
_own_union
;
[|
set_solver
..].
rewrite
[
E
as
X
in
na
_own
tid
X
](
union_difference_L
(
↑
N
)
E
)
//.
rewrite
[
X
in
(
X
∪
_
)](
union_difference_L
{[
i
]}
(
↑
N
))
?
na
_own_union
;
[|
set_solver
..].
iDestruct
"Htoks"
as
"[[Htoki $] $]"
.
iInv
N
as
"[[$ >Hdis]|>Htoki2]"
"Hclose"
.
-
iMod
(
"Hclose"
with
"[Htoki]"
)
as
"_"
;
first
auto
.
...
...
@@ -86,6 +88,6 @@ Section proofs.
+
iDestruct
(
own_valid_2
with
"Hdis Hdis2"
)
as
%[
_
Hval
%
gset_disj_valid_op
].
set_solver
.
+
iFrame
.
iApply
"Hclose"
.
iNext
.
iLeft
.
by
iFrame
.
-
iDestruct
(
tl
_own_disjoint
with
"Htoki Htoki2"
)
as
%?.
set_solver
.
-
iDestruct
(
na
_own_disjoint
with
"Htoki Htoki2"
)
as
%?.
set_solver
.
Qed
.
End
proofs
.
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new 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