Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
7
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Simon Friis Vindum
examples
Commits
467e6866
Commit
467e6866
authored
Apr 22, 2020
by
Simon Friis Vindum
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Fix the last admits in the refinement proof
parent
5d088fe1
Changes
1
Show whitespace changes
Inline
Side-by-side
Showing
1 changed file
with
12 additions
and
29 deletions
+12
-29
theories/logrel/F_mu_ref_conc/examples/queue/refinement.v
theories/logrel/F_mu_ref_conc/examples/queue/refinement.v
+12
-29
No files found.
theories/logrel/F_mu_ref_conc/examples/queue/refinement.v
View file @
467e6866
...
...
@@ -27,22 +27,6 @@ Canonical Structure gnameO := leibnizO gname.
Definition
mapUR
:
ucmraT
:
=
gmapUR
loc
(
agreeR
(
leibnizO
(
gname
*
loc
))).
Definition
nodeUR
:
ucmraT
:
=
authUR
(
gmapUR
loc
(
agreeR
(
leibnizO
(
gname
*
loc
)))).
Section
my_map
.
Context
`
{
Monoid
M
o
}.
Context
`
{
Countable
K
}
{
A
:
Type
}.
Implicit
Types
m
:
gmap
K
A
.
Context
{
PROP
:
bi
}.
Implicit
Types
Φ
:
K
→
A
→
PROP
.
Infix
"`o`"
:
=
o
(
at
level
50
,
left
associativity
).
(* This lemma is committed upstream and should be available from Iris at
some point (when?) *)
Lemma
big_sepM_insert_delete
Φ
m
i
x
:
(([
∗
map
]
k
↦
y
∈
<[
i
:
=
x
]>
m
,
Φ
k
y
)
⊣
⊢
Φ
i
x
∗
[
∗
map
]
k
↦
y
∈
(
delete
i
m
),
Φ
k
y
)%
I
.
Proof
.
Admitted
.
End
my_map
.
Section
Queue_refinement
.
Context
`
{
heapIG
Σ
,
cfgSG
Σ
,
inG
Σ
fracAgreeR
,
inG
Σ
exlTokR
,
inG
Σ
nodeUR
,
inG
Σ
nodeStateR
}.
...
...
@@ -235,10 +219,10 @@ Section Queue_refinement.
Lemma
map_singleton_included
(
m
:
gmap
loc
(
gname
*
loc
))
(
l
:
loc
)
v
:
({[
l
:
=
to_agree
v
]}
:
mapUR
)
≼
((
to_agree
<$>
m
)
:
mapUR
)
→
m
!!
l
=
Some
v
.
Proof
.
rewrit
e
singleton_included_l
=>
-[
[
q'
av
]
[]
].
(*
rewrite lookup_fmap fmap_Some_equiv => -[
v' [Hl [/=
-> ->]]].
move
=>
/Some_
pair_
included_total
_2 [_]
/to_agree_included /leibniz_equiv_iff ->
//. *)
Admitt
ed
.
mov
e
/
singleton_included_l
=>
-[
y
].
rewrite
lookup_fmap
fmap_Some_equiv
=>
-[
[
x
[
->
->]]].
by
move
/
Some_included_total
/
to_agree_included
/
leibniz_equiv_iff
->
.
Q
ed
.
Lemma
auth_node_mapsto_Some
γ
m
ℓ
s
ι
ℓ
n
:
own
γ
(
●
(
to_agree
<$>
m
)
:
nodeUR
)
-
∗
...
...
@@ -256,7 +240,7 @@ Section Queue_refinement.
iDestruct
(
own_valid_2
with
"a b"
)
as
%
Hv
.
rewrite
-
auth_frag_op
in
Hv
.
apply
(
auth_frag_valid
(
_
⋅
_
))
in
Hv
.
(* Why is this necessary? *)
rewrite
op_
singleton
in
Hv
.
rewrite
singleton
_op
in
Hv
.
apply
singleton_valid
,
agree_op_invL'
in
Hv
.
inversion
Hv
.
done
.
...
...
@@ -274,8 +258,7 @@ Section Queue_refinement.
iDestruct
(
big_sepM_insert_delete
with
"[new $mon]"
)
as
"mon"
.
{
iExists
ι
,
ℓ
n
.
iFrame
.
done
.
}
rewrite
insert_id
;
done
.
(* FIXME: Coq complains: Some unresolved existential variables remain *)
Admitted
.
Qed
.
(* Reinsert a node that has been taken out. *)
Lemma
reinsert_node
γ
κ
ℓ
q
ℓ
s
ι
ℓ
n
(
m
:
gmap
loc
(
gname
*
loc
))
:
...
...
@@ -381,16 +364,16 @@ Section Queue_refinement.
iFrame
.
iExists
_
,
_
.
iFrame
"tailPts"
.
iFrame
.
-
iDestruct
"nodeList"
as
(
ℓ
0
ℓ
next
q
ι
)
"(FOO & BAR & BAZ & HIHI & nodeListTail)"
.
iApply
(
"IH"
with
"authM bigSep nilPts tailPts nodePts nodeListTail ℓPts tok"
).
iNext
.
iIntros
"HI"
.
iNext
.
iDestruct
1
as
(
ι
'
)
"(authM & bigSep & nilPts & tailPts & nodePts & nodeListTail & ℓPts)"
.
iApply
"Hϕ"
.
iDestruct
"HI"
as
(
ι
'
)
"(authM & bigSep & nilPts & tailPts & nodePts & nodeListTail & ℓPts)"
.
iExistsFrame
.
Qed
.
(* This lemma has been commited upstream to Iris and will be available in the future. *)
Lemma
auth_update_core_id_frac
{
A
:
ucmraT
}
(
a
b
:
A
)
`
{!
CoreId
b
}
q
:
b
≼
a
→
●
{
q
}
a
~~>
●
{
q
}
a
⋅
◯
b
.
Proof
.
Admitted
.
Proof
.
Admitted
.
Lemma
MS_CG_counter_refinement
:
[]
⊨
MS_queue
≤
log
≤
CG_queue
:
...
...
@@ -808,8 +791,8 @@ Section Queue_refinement.
simpl
.
iApply
(
enqueue_cas
with
"[$authM $bigSep $nodesInv $nilPts $tailPts $nodePts $nodeList $ℓPts $tok]"
).
{
done
.
}
iNext
.
iIntros
"HI"
.
iDestruct
"HI"
as
(
ι
3
)
"(authM & bigSep & nodePts & nodeList & #frag & ℓPts)"
.
iNext
.
iDestruct
1
as
(
ι
3
)
"(authM & bigSep & nodePts & nodeList & #frag & ℓPts)"
.
iMod
(
steps_CG_enqueue
with
"[$Hspec $Hj $lofal $Hsq]"
)
as
"(Hj & Hsq & lofal)"
.
{
solve_ndisj
.
}
iDestruct
(
reinsert_node
with
"authM fragO bigSep [sentinelPts ℓPts nodePts rest]"
)
as
"(authM & bigSep)"
.
...
...
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