Skip to content
GitLab
Menu
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
FP
iris-atomic
Commits
889d57f1
Commit
889d57f1
authored
Sep 22, 2016
by
Zhen Zhang
Browse files
change invariant
parent
d5272bc5
Changes
1
Hide whitespace changes
Inline
Side-by-side
srv.v
View file @
889d57f1
...
@@ -3,7 +3,7 @@ From iris.proofmode Require Import invariants ghost_ownership.
...
@@ -3,7 +3,7 @@ From iris.proofmode Require Import invariants ghost_ownership.
From
iris
.
heap_lang
Require
Export
lang
.
From
iris
.
heap_lang
Require
Export
lang
.
From
iris
.
heap_lang
Require
Import
proofmode
notation
.
From
iris
.
heap_lang
Require
Import
proofmode
notation
.
From
iris
.
heap_lang
.
lib
Require
Import
spin_lock
.
From
iris
.
heap_lang
.
lib
Require
Import
spin_lock
.
From
iris
.
algebra
Require
Import
frac
excl
dec_agree
upred_big_op
gset
gmap
.
From
iris
.
algebra
Require
Import
upred
frac
excl
dec_agree
upred_big_op
gset
gmap
.
From
iris
.
tests
Require
Import
atomic
treiber_stack
.
From
iris
.
tests
Require
Import
atomic
treiber_stack
.
From
flatcomb
Require
Import
misc
.
From
flatcomb
Require
Import
misc
.
...
@@ -41,69 +41,39 @@ Definition flat : val :=
...
@@ -41,69 +41,39 @@ Definition flat : val :=
Global
Opaque
doOp
install
loop
flat
.
Global
Opaque
doOp
install
loop
flat
.
Definition
hdset
:=
gset
loc
.
Definition
gnmap
:=
gmap
loc
(
dec_agree
(
gname
*
gname
*
gname
*
gname
)).
Definition
srvR
:=
prodR
fracR
(
dec_agreeR
val
).
Definition
srvR
:=
prodR
fracR
(
dec_agreeR
val
).
Definition
hdsetR
:=
gset_disjUR
loc
.
Class
srvG
Σ
:=
SrvG
{
srv_G
:>
inG
Σ
srvR
}
.
Definition
gnmapR
:=
gmapUR
loc
(
dec_agreeR
(
gname
*
gname
*
gname
*
gname
)).
Definition
srv
Σ
:
gFunctors
:=
#[
GFunctor
(
constRF
srvR
)].
Class
srvG
Σ
:=
SrvG
{
srv_tokG
:>
inG
Σ
srvR
;
hd_G
:>
inG
Σ
(
authR
hdsetR
);
gn_G
:>
inG
Σ
(
authR
gnmapR
)
}
.
Definition
srv
Σ
:
gFunctors
:=
#[
GFunctor
(
constRF
srvR
);
GFunctor
(
constRF
(
authR
hdsetR
));
GFunctor
(
constRF
(
authR
gnmapR
))
].
Instance
subG_srv
Σ
{
Σ
}
:
subG
srv
Σ
Σ
→
srvG
Σ
.
Instance
subG_srv
Σ
{
Σ
}
:
subG
srv
Σ
Σ
→
srvG
Σ
.
Proof
.
intros
[
?%
subG_inG
[
?
subG_inG
[
?
subG_inG
_
]
%
subG_inv
]
%
subG_inv
]
%
subG_inv
.
split
;
apply
_.
Qed
.
Proof
.
intros
[
?%
subG_inG
_
]
%
subG_inv
.
split
;
apply
_.
Qed
.
Section
proof
.
Section
proof
.
Context
`
{!
heapG
Σ
,
!
lockG
Σ
,
!
srvG
Σ
}
(
N
:
namespace
).
Context
`
{!
heapG
Σ
,
!
lockG
Σ
,
!
evidenceG
loc
val
Σ
,
!
srvG
Σ
}
(
N
:
namespace
).
Definition
p_inv
Definition
p_inv
(
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
(
p
:
loc
)
(
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
(
Q
:
val
→
val
→
Prop
)
:
iProp
Σ
:=
(
Q
:
val
→
val
→
Prop
)
((
∃
(
y
:
val
),
p
↦
InjRV
y
★
own
γ
1
(
Excl
())
★
own
γ
3
(
Excl
()))
∨
(
v
:
val
)
:
iProp
Σ
:=
(
∃
(
x
:
val
),
p
↦
InjLV
x
★
own
γ
x
((
1
/
2
)
%
Qp
,
DecAgree
x
)
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
()))
∨
((
∃
(
y
:
val
),
v
=
InjRV
y
★
own
γ
1
(
Excl
())
★
own
γ
3
(
Excl
()))
∨
(
∃
(
x
:
val
),
p
↦
InjLV
x
★
own
γ
x
((
1
/
4
)
%
Qp
,
DecAgree
x
)
★
own
γ
2
(
Excl
())
★
own
γ
4
(
Excl
()))
∨
(
∃
(
x
:
val
),
v
=
InjLV
x
★
own
γ
x
((
1
/
2
)
%
Qp
,
DecAgree
x
)
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
()))
∨
(
∃
(
x
y
:
val
),
p
↦
InjRV
y
★
own
γ
x
((
1
/
2
)
%
Qp
,
DecAgree
x
)
★
■
Q
x
y
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
())))
%
I
.
(
∃
(
x
:
val
),
v
=
InjLV
x
★
own
γ
x
((
1
/
4
)
%
Qp
,
DecAgree
x
)
★
own
γ
2
(
Excl
())
★
own
γ
4
(
Excl
()))
∨
(
∃
(
x
y
:
val
),
v
=
InjRV
y
★
own
γ
x
((
1
/
2
)
%
Qp
,
DecAgree
x
)
★
■
Q
x
y
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
())))
%
I
.
Definition
p_inv
'
γ
2
(
γ
s
:
dec_agree
(
gname
*
gname
*
gname
*
gname
))
p
Q
:=
Definition
p_inv_R
γ
2
Q
v
:
iProp
Σ
:=
(
∃
γ
x
γ
1
γ
3
γ
4
:
gname
,
p_inv
γ
x
γ
1
γ
2
γ
3
γ
4
Q
v
)
%
I
.
match
γ
s
with
|
DecAgreeBot
=>
False
%
I
|
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
)
=>
p_inv
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
end
.
Definition
srv_inv
(
γ
hd
γ
gn
γ
2
:
gname
)
(
s
:
loc
)
(
Q
:
val
→
val
→
Prop
)
:
iProp
Σ
:=
Definition
srv_inv
(
γ
γ
2
:
gname
)
(
s
:
loc
)
(
Q
:
val
→
val
→
Prop
)
:
iProp
Σ
:=
(
∃
(
hds
:
hdset
)
(
gnm
:
gnmap
),
(
∃
xs
,
is_stack
'
(
p_inv_R
γ
2
Q
)
xs
s
γ
)
%
I
.
own
γ
hd
(
●
GSet
hds
)
★
own
γ
gn
(
●
gnm
)
★
(
∃
xs
:
list
loc
,
is_stack
s
(
map
(
fun
x
=>
#
(
LitLoc
x
))
xs
)
★
[
★
list
]
k
↦
x
∈
xs
,
■
(
x
∈
dom
(
gset
loc
)
gnm
))
★
([
★
set
]
hd
∈
hds
,
∃
xs
,
is_list
hd
(
map
(
fun
x
=>
#
(
LitLoc
x
))
xs
)
★
[
★
list
]
k
↦
x
∈
xs
,
■
(
x
∈
dom
(
gset
loc
)
gnm
))
★
([
★
map
]
p
↦
γ
s
∈
gnm
,
p_inv
'
γ
2
γ
s
p
Q
)
)
%
I
.
Instance
p_inv_timeless
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
:
TimelessP
(
p_inv
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
).
Instance
p_inv_timeless
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
:
TimelessP
(
p_inv
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
).
Proof
.
apply
_.
Qed
.
Proof
.
apply
_.
Qed
.
Instance
p
_inv
'
_
timeless
γ
2
γ
s
p
Q
:
TimelessP
(
p
_inv
'
γ
2
γ
s
p
Q
).
Instance
srv
_inv_timeless
γ
γ
2
s
Q
:
TimelessP
(
srv
_inv
γ
γ
2
s
Q
).
Proof
.
Proof
.
rewrite
/
p_inv
'
.
destruct
γ
s
as
[
γ
s
|
]
.
apply
uPred
.
exist_timeless
.
-
repeat
(
destruct
γ
s
as
[
γ
s
?
]).
apply
_
.
move
=>
x
.
apply
is_stack
'_
timeless
.
-
apply
_.
move
=>
v
.
apply
_.
Qed
.
Qed
.
Instance
srv_inv_timeless
γ
hd
γ
gn
γ
2
s
Q
:
TimelessP
(
srv_inv
γ
hd
γ
gn
γ
2
s
Q
).
Proof
.
apply
_.
Qed
.
(
*
Lemma
push_spec
*
)
(
*
Lemma
push_spec
*
)
(
*
(
Φ
:
val
→
iProp
Σ
)
(
Q
:
val
→
val
→
Prop
)
*
)
(
*
(
Φ
:
val
→
iProp
Σ
)
(
Q
:
val
→
val
→
Prop
)
*
)
(
*
(
p
:
loc
)
(
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
*
)
(
*
(
p
:
loc
)
(
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
*
)
...
@@ -237,7 +207,7 @@ Section proof.
...
@@ -237,7 +207,7 @@ Section proof.
Lemma
loop_iter_list_spec
Φ
(
f
:
val
)
(
s
hd
:
loc
)
Q
(
γ
hd
γ
gn
γ
2
:
gname
)
xs
:
Lemma
loop_iter_list_spec
Φ
(
f
:
val
)
(
s
hd
:
loc
)
Q
(
γ
hd
γ
gn
γ
2
:
gname
)
xs
:
heapN
⊥
N
→
heapN
⊥
N
→
heap_ctx
★
inv
N
(
srv_inv
γ
hd
γ
gn
γ
2
s
Q
)
★
□
(
∀
x
:
val
,
WP
f
x
{{
v
,
■
Q
x
v
}}
)
%
I
★
own
γ
2
(
Excl
())
★
heap_ctx
★
inv
N
(
srv_inv
γ
hd
γ
gn
γ
2
s
Q
)
★
□
(
∀
x
:
val
,
WP
f
x
{{
v
,
■
Q
x
v
}}
)
%
I
★
own
γ
2
(
Excl
())
★
is_list
hd
xs
★
own
γ
hd
(
◯
GSet
{
[
hd
]
}
)
★
(
own
γ
2
(
Excl
())
-
★
Φ
#())
is_list
hd
xs
★
own
γ
hd
(
◯
(
{
[
hd
]
}
:
hdsetR
)
)
★
(
own
γ
2
(
Excl
())
-
★
Φ
#())
⊢
WP
doOp
f
{{
f
'
,
WP
iter
'
#
hd
f
'
{{
Φ
}}
}}
.
⊢
WP
doOp
f
{{
f
'
,
WP
iter
'
#
hd
f
'
{{
Φ
}}
}}
.
Proof
.
Proof
.
iIntros
(
HN
)
"(#Hh & #? & #Hf & Hxs & Hhd & Ho2 & HΦ)"
.
iIntros
(
HN
)
"(#Hh & #? & #Hf & Hxs & Hhd & Ho2 & HΦ)"
.
...
@@ -287,7 +257,7 @@ Section proof.
...
@@ -287,7 +257,7 @@ Section proof.
iAssert
(
∃
(
hd
'0
:
loc
)
(
q0
:
Qp
),
iAssert
(
∃
(
hd
'0
:
loc
)
(
q0
:
Qp
),
hd
↦
{
q0
}
SOMEV
(#
p
,
#
hd
'0
)
★
is_list
hd
'0
xs
'
)
%
I
with
"[Hhd2 Hl]"
as
"He"
.
hd
↦
{
q0
}
SOMEV
(#
p
,
#
hd
'0
)
★
is_list
hd
'0
xs
'
)
%
I
with
"[Hhd2 Hl]"
as
"He"
.
{
iExists
hd
'
,
(
q
/
2
)
%
Qp
.
by
iFrame
.
}
{
iExists
hd
'
,
(
q
/
2
)
%
Qp
.
by
iFrame
.
}
iAssert
(
own
γ
hd
(
◯
GSet
{
[
hd
]
}
))
as
"Hfrag"
.
iAssert
(
own
γ
hd
(
◯
{
[
hd
]
}
))
as
"Hfrag"
.
{
admit
.
}
{
admit
.
}
iSpecialize
(
"IH"
with
"Ho2 He Hfrag HΦ"
).
iSpecialize
(
"IH"
with
"Ho2 He Hfrag HΦ"
).
admit
.
admit
.
...
@@ -304,7 +274,7 @@ Section proof.
...
@@ -304,7 +274,7 @@ Section proof.
⊢
WP
iter
(
doOp
f
)
#
s
{{
Φ
}}
.
⊢
WP
iter
(
doOp
f
)
#
s
{{
Φ
}}
.
Proof
.
Proof
.
iIntros
(
HN
)
"(#Hh & #? & #? & ? & ?)"
.
iIntros
(
HN
)
"(#Hh & #? & #? & ? & ?)"
.
iAssert
(
∃
(
hd
:
loc
)
xs
,
is_list
hd
xs
★
own
γ
hd
(
◯
GSet
{
[
hd
]
}
)
★
s
↦
#
hd
)
%
I
as
"H"
.
iAssert
(
∃
(
hd
:
loc
)
xs
,
is_list
hd
xs
★
own
γ
hd
(
◯
{
[
hd
]
}
)
★
s
↦
#
hd
)
%
I
as
"H"
.
{
admit
.
}
{
admit
.
}
iDestruct
"H"
as
(
hd
xs
)
"(? & ? & ?)"
.
iDestruct
"H"
as
(
hd
xs
)
"(? & ? & ?)"
.
wp_bind
(
doOp
_
).
wp_bind
(
doOp
_
).
...
@@ -343,19 +313,23 @@ Section proof.
...
@@ -343,19 +313,23 @@ Section proof.
Lemma
flat_spec
(
f
:
val
)
Q
:
Lemma
flat_spec
(
f
:
val
)
Q
:
heapN
⊥
N
→
heapN
⊥
N
→
heap_ctx
★
□
(
∀
x
:
val
,
WP
f
x
{{
v
,
■
Q
x
v
}}
)
%
I
heap_ctx
★
□
(
∀
x
:
val
,
WP
f
x
{{
v
,
■
Q
x
v
}}
)
%
I
⊢
WP
flat
f
{{
f
'
,
True
%
I
}}
.
⊢
WP
flat
f
{{
f
'
,
□
(
∀
x
:
val
,
WP
f
'
x
{{
v
,
■
Q
x
v
}}
)
}}
.
Proof
.
Proof
.
iIntros
(
HN
)
"(#Hh & #?)"
.
iIntros
(
HN
)
"(#Hh & #?)"
.
wp_seq
.
wp_alloc
lk
as
"Hl"
.
wp_seq
.
wp_alloc
lk
as
"Hl"
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
2
)
"Ho2"
;
first
done
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
2
)
"Ho2"
;
first
done
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
lk
)
"Hγlk"
;
first
done
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
lk
)
"Hγlk"
;
first
done
.
iVs
(
own_alloc
(
●
(
∅
:
hdsetR
)
⋅
◯
∅
))
as
(
γ
hd
)
"[Hgs Hgs']"
;
first
admit
.
iVs
(
own_alloc
(
●
∅
⋅
◯
∅
))
as
(
γ
gn
)
"[Hgs Hgs']"
;
first
admit
.
iVs
(
own_alloc
())
as
(
γ
lk
)
"Hγlk"
;
first
done
.
iVs
(
inv_alloc
N
_
(
lock_inv
γ
lk
lk
(
own
γ
2
(
Excl
())))
with
"[-]"
)
as
"#?"
.
iVs
(
inv_alloc
N
_
(
lock_inv
γ
lk
lk
(
own
γ
2
(
Excl
())))
with
"[-]"
)
as
"#?"
.
{
iIntros
"!>"
.
iExists
false
.
by
iFrame
.
}
{
iIntros
"!>"
.
iExists
false
.
by
iFrame
.
}
wp_let
.
wp_bind
(
new_stack
_
).
wp_let
.
wp_bind
(
new_stack
_
).
iApply
new_stack_spec
=>
//.
iApply
new_stack_spec
=>
//.
iFrame
"Hh"
.
iIntros
(
s
)
"Hs"
.
iFrame
"Hh"
.
iIntros
(
s
)
"Hs"
.
iVs
(
inv_alloc
N
_
(
srv_inv
γ
hd
γ
gn
γ
2
s
Q
)
with
""
)
as
"#?"
.
wp_let
.
wp_let
.
done
.
iVsIntro
.
iPureIntro
Qed
.
Qed
.
End
proof
.
\ No newline at end of file
Write
Preview
Supports
Markdown
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