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
Dan Frumin
iris-atomic
Commits
535a2fc5
Commit
535a2fc5
authored
Dec 16, 2016
by
Ralf Jung
Browse files
FAILED: port evmap to dec_agree
parent
a50a88fc
Changes
4
Hide whitespace changes
Inline
Side-by-side
opam.pins
View file @
535a2fc5
coq-iris https://gitlab.mpi-sws.org/FP/iris-coq
766dbcd2415df9256f197dc562a0a15f9b0ddeac
coq-iris https://gitlab.mpi-sws.org/FP/iris-coq
a953a68d9b9f30f8aa4e0e36811b9175f3f2ea58
theories/evmap.v
View file @
535a2fc5
(
*
evmap
.
v
--
generalized
heap
-
like
monoid
composite
*
)
From
iris
.
base_logic
Require
Export
invariants
.
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris
.
algebra
Require
Export
auth
frac
gmap
.
From
iris
.
algebra
Require
deprecated
.
From
iris
.
algebra
Require
Export
auth
frac
agree
gmap
.
From
iris
.
proofmode
Require
Import
tactics
.
Export
deprecated
.
dec_agree
.
Section
evmap
.
Context
(
K
A
:
Type
)
(
Q
:
cmraT
)
`
{
Countable
K
,
EqDecision
A
}
.
Definition
evkR
:=
prodR
Q
(
dec_agreeR
A
).
Canonical
AC
:=
leibnizC
A
.
Definition
evkR
:=
prodR
Q
(
agreeR
AC
).
Definition
evmapR
:=
gmapUR
K
evkR
.
Definition
evidenceR
:=
authR
evmapR
.
Class
evidenceG
Σ
:=
EvidenceG
{
evidence_G
:>
inG
Σ
evidenceR
}
.
...
...
@@ -20,32 +18,31 @@ Section evmap.
(
*
Some
basic
supporting
lemmas
*
)
Lemma
map_agree_eq
m
m
'
(
hd
:
K
)
(
p
q
:
Q
)
(
x
y
:
A
)
:
m
!!
hd
=
Some
(
p
,
DecA
gree
y
)
→
m
=
{
[
hd
:=
(
q
,
DecA
gree
x
)]
}
⋅
m
'
→
x
=
y
.
m
!!
hd
≡
Some
(
p
,
to_a
gree
y
)
→
m
≡
{
[
hd
:=
(
q
,
to_a
gree
x
)]
}
⋅
m
'
→
x
=
y
.
Proof
.
intros
H1
H2
.
destruct
(
decide
(
x
=
y
))
as
[
->|
Hneq
]
=>
//.
exfalso
.
subst
.
rewrite
lookup_op
lookup_singleton
in
H1
.
destruct
(
m
'
!!
hd
)
as
[[
b
[
c
|
]]
|
]
eqn
:
Heqn
;
rewrite
Heqn
in
H1
;
inversion
H1
=>
//.
destruct
(
decide
(
x
=
c
))
as
[
->|
Hneq3
]
=>
//.
-
rewrite
dec_agree_idemp
in
H3
.
by
inversion
H3
.
-
rewrite
dec_agree_ne
in
H3
=>
//.
intros
H1
H2
.
unfold_leibniz
.
eapply
to_agree_included
.
assert
(
{
[
hd
:=
(
q
,
to_agree
x
)]
}
!!
hd
≼
m
!!
hd
)
as
Hincl
.
{
apply
lookup_included
.
exists
m
'
.
done
.
}
move:
Hincl
.
rewrite
lookup_singleton
H1
.
move
/
Some_included
=>
[[
/=
_
->
]
|/
prod_included
[
_
?
]]
//.
Qed
.
(
*
Lemma
map_agree_somebot
m
m
'
(
hd
:
K
)
(
p
q
:
Q
)
(
x
:
A
)
:
m
!!
hd
=
Some
(
p
,
DecAgreeBot
)
→
m
'
!!
hd
=
None
→
m
=
{
[
hd
:=
(
q
,
DecA
gree
x
)]
}
⋅
m
'
→
False
.
m
=
{
[
hd
:=
(
q
,
to_a
gree
x
)]
}
⋅
m
'
→
False
.
Proof
.
intros
H1
H2
H3
.
subst
.
rewrite
lookup_op
lookup_singleton
in
H1
.
destruct
(
m
'
!!
hd
)
as
[[
b
[
c
|
]]
|
]
eqn
:
Heqn
;
rewrite
Heqn
in
H1
;
inversion
H1
=>
//.
Qed
.
*
)
Lemma
map_agree_none
m
m
'
(
hd
:
K
)
(
q
:
Q
)
(
x
:
A
)
:
m
!!
hd
=
None
→
m
=
{
[
hd
:=
(
q
,
DecA
gree
x
)]
}
⋅
m
'
→
False
.
m
!!
hd
=
None
→
m
≡
{
[
hd
:=
(
q
,
to_a
gree
x
)]
}
⋅
m
'
→
False
.
Proof
.
intros
H1
H2
.
subst
.
rewrite
lookup_op
lookup_singleton
in
H1
.
destruct
(
m
'
!!
hd
)
=>
//
.
intros
H1
H2
.
move
:
(
H2
hd
)
.
rewrite
H1
lookup_op
lookup_singleton
.
case:
(
m
'
!!
hd
)
=>
[
?|
]
//=; intros XX; inversion XX
.
Qed
.
End
evmap
.
...
...
@@ -54,7 +51,7 @@ Section evmapR.
Context
`
{!
inG
Σ
(
authR
(
evmapR
K
A
unitR
))
}
.
(
*
Evidence
that
k
immutably
maps
to
some
fixed
v
*
)
Definition
ev
γ
m
(
k
:
K
)
(
v
:
A
)
:=
own
γ
m
(
◯
{
[
k
:=
((),
DecA
gree
v
)
]
}
)
%
I
.
Definition
ev
γ
m
(
k
:
K
)
(
v
:
A
)
:=
own
γ
m
(
◯
{
[
k
:=
((),
to_a
gree
v
)
]
}
)
%
I
.
Global
Instance
persistent_ev
γ
m
k
v
:
PersistentP
(
ev
γ
m
k
v
).
Proof
.
apply
_.
Qed
.
...
...
@@ -62,7 +59,7 @@ Section evmapR.
(
*
Alloc
a
new
mapsto
*
)
Lemma
evmap_alloc
γ
m
m
k
v
:
m
!!
k
=
None
→
own
γ
m
(
●
m
)
⊢
|==>
own
γ
m
(
●
(
<
[
k
:=
((),
DecA
gree
v
)
]
>
m
)
⋅
◯
{
[
k
:=
((),
DecA
gree
v
)
]
}
).
own
γ
m
(
●
m
)
⊢
|==>
own
γ
m
(
●
(
<
[
k
:=
((),
to_a
gree
v
)
]
>
m
)
⋅
◯
{
[
k
:=
((),
to_a
gree
v
)
]
}
).
Proof
.
iIntros
(
?
)
"Hm"
.
iDestruct
(
own_update
with
"Hm"
)
as
"?"
;
last
by
iAssumption
.
...
...
@@ -77,36 +74,34 @@ Section evmapR.
iIntros
(
?
)
"[Hom Hev]"
.
iCombine
"Hom"
"Hev"
as
"Hauth"
.
iDestruct
(
own_valid
with
"Hauth"
)
as
%
Hvalid
.
iPureIntro
.
apply
auth_valid_discrete_2
in
Hvalid
as
[[
af
Compose
%
leibniz_equiv_iff
]
Valid
].
eapply
(
map_agree_none
_
_
_
m
)
=>
//.
apply
auth_valid_discrete_2
in
Hvalid
as
[[
af
Compose
]
Valid
].
eapply
(
map_agree_none
_
_
_
m
)
=>
//.
Qed
.
Lemma
map_agree_eq
'
γ
m
m
hd
x
agy
:
m
!!
hd
=
Some
((),
agy
)
→
ev
γ
m
hd
x
∗
own
γ
m
(
●
m
)
⊢
⌜
DecA
gree
x
=
agy
⌝
.
ev
γ
m
hd
x
∗
own
γ
m
(
●
m
)
⊢
⌜
to_a
gree
x
≡
agy
⌝
.
Proof
.
iIntros
(
?
)
"[#Hev Hom]"
.
iCombine
"Hom"
"Hev"
as
"Hauth"
.
iDestruct
(
own_valid
γ
m
(
●
m
⋅
◯
{
[
hd
:=
(
_
,
DecA
gree
x
)]
}
)
iDestruct
(
own_valid
γ
m
(
●
m
⋅
◯
{
[
hd
:=
(
_
,
to_a
gree
x
)]
}
)
with
"[Hauth]"
)
as
%
Hvalid
=>
//.
apply
auth_valid_discrete_2
in
Hvalid
as
[[
af
Compose
%
leibniz_equiv_iff
]
Valid
].
destruct
agy
as
[
y
|
].
-
iDestruct
"Hauth"
as
"[? ?]"
.
iFrame
.
iPureIntro
.
apply
f_equal
.
eapply
(
map_agree_eq
_
_
_
m
)
=>
//.
-
iAssert
(
✓
m
)
%
I
as
"H"
=>
//. rewrite (gmap_validI m).
iDestruct
(
"H"
$
!
hd
)
as
"%"
.
exfalso
.
subst
.
rewrite
H0
in
H1
.
by
destruct
H1
as
[
?
?
].
apply
auth_valid_discrete_2
in
Hvalid
as
[[
af
Compose
]
Valid
].
iDestruct
"Hauth"
as
"[? ?]"
.
iPureIntro
.
edestruct
(
to_agree_uninj
agy
)
as
[
y
Heq
].
-
move
:
(
Valid
hd
).
rewrite
H0
=>-
[
??
]
//.
-
rewrite
-
Heq
.
f_equiv
.
eapply
(
map_agree_eq
_
_
_
m
)
=>
//. by rewrite Heq H0.
Qed
.
(
*
Evidence
is
the
witness
of
membership
*
)
Lemma
ev_map_witness
γ
m
m
hd
x
:
ev
γ
m
hd
x
∗
own
γ
m
(
●
m
)
⊢
⌜
m
!!
hd
=
Some
(
∅
,
DecA
gree
x
)
⌝
.
ev
γ
m
hd
x
∗
own
γ
m
(
●
m
)
⊢
⌜
m
!!
hd
≡
Some
(
∅
,
to_a
gree
x
)
⌝
.
Proof
.
iIntros
"[#Hev Hom]"
.
destruct
(
m
!!
hd
)
as
[[[]
agy
]
|
]
eqn
:
Heqn
.
-
iDestruct
(
map_agree_eq
'
with
"[-]"
)
as
%
H
'
=>
//; first by iFrame. by subst.
-
iDestruct
(
map_agree_eq
'
with
"[-]"
)
as
%
H
'
=>
//; first by iFrame.
iPureIntro
.
by
rewrite
H
'
.
-
iExFalso
.
iApply
map_agree_none
'
=>
//. by iFrame.
Qed
.
...
...
@@ -118,12 +113,12 @@ Section evmapR.
destruct
(
decide
(
a1
=
a2
))
as
[
->|
Hneq
].
-
by
iFrame
.
-
iCombine
"Ho"
"Ho'"
as
"Ho"
.
rewrite
-
(
@
auth_frag_op
(
evmapR
K
A
unitR
)
{
[
p
:=
(
_
,
DecA
gree
a1
)]
}
{
[
p
:=
(
_
,
DecA
gree
a2
)]
}
).
rewrite
-
(
@
auth_frag_op
(
evmapR
K
A
unitR
)
{
[
p
:=
(
_
,
to_a
gree
a1
)]
}
{
[
p
:=
(
_
,
to_a
gree
a2
)]
}
).
iDestruct
(
own_valid
with
"Ho"
)
as
%
Hvalid
.
exfalso
.
rewrite
op_singleton
in
Hvalid
.
apply
auth_valid_discrete
in
Hvalid
.
simpl
in
Hvalid
.
apply
singleton_valid
in
Hvalid
.
destruct
Hvalid
as
[
_
Hvalid
].
apply
dec_
agree_op_inv
in
Hvalid
.
inversion
Hvalid
.
subst
.
auto
.
apply
agree_op_inv
in
Hvalid
.
apply
(
inj
to_agree
)
in
Hvalid
.
by
fold_leibniz
.
Qed
.
End
evmapR
.
theories/flat.v
View file @
535a2fc5
...
...
@@ -158,7 +158,7 @@ Section proof.
iDestruct
(
evmap_alloc
_
_
_
m
p
(
γ
x
,
γ
1
,
γ
3
,
γ
4
,
γ
q
)
with
"[Hm]"
)
as
">[Hm1 #Hm2]"
=>
//.
iDestruct
"Hl"
as
"[Hl1 Hl2]"
.
iMod
(
"Hclose"
with
"[HRm Hm1 Hl1 Hrs]"
).
+
iNext
.
iFrame
.
iExists
(
<
[
p
:=
(
∅
,
DecA
gree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
,
γ
q
))]
>
m
).
iFrame
.
+
iNext
.
iFrame
.
iExists
(
<
[
p
:=
(
∅
,
to_a
gree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
,
γ
q
))]
>
m
).
iFrame
.
iDestruct
(
big_sepM_insert
_
m
with
"[-]"
)
as
"H"
=>
//.
iSplitL
"Hl1"
;
last
by
iAssumption
.
eauto
.
+
iDestruct
(
own_update
with
"Hx"
)
as
">[Hx1 Hx2]"
;
first
by
apply
pair_l_frac_op_1
'
.
...
...
@@ -193,8 +193,12 @@ Section proof.
iDestruct
"HRs"
as
(
m
)
"[>Hom HRs]"
.
(
*
acccess
*
)
iDestruct
(
ev_map_witness
_
_
_
m
with
"[Hev Hom]"
)
as
%
H
'
;
first
by
iFrame
.
case
Hlk
:
(
m
!!
xhd
)
=>
[
xx
|
];
last
first
.
{
rewrite
Hlk
in
H
'
.
inversion
H
'
.
}
iDestruct
(
big_sepM_delete_later
(
perR
_
)
m
with
"HRs"
)
as
"[Hp HRm]"
=>
//.
iDestruct
"Hp"
as
(
v
'
)
"[>% [Hpinv' >Hahd]]"
.
inversion
H
.
subst
.
rewrite
Hlk
in
H
'
.
iDestruct
"Hp"
as
(
v
'
)
"[>% [Hpinv' >Hahd]]"
.
inversion
H
'
.
subst
.
destruct
H2
as
[
_
H2
].
apply
(
inj
to_agree
)
in
H2
.
fold_leibniz
.
subst
.
iDestruct
"Hpinv'"
as
(
ts
p
''
)
"[>% [>#Hevm [Hp | [Hp | [Hp | Hp]]]]]"
;
subst
.
+
iDestruct
"Hp"
as
(
y
)
"(>Hp & Hs)"
.
wp_load
.
iMod
(
"Hclose"
with
"[-Hor HR Hev HΦ']"
).
...
...
@@ -228,10 +232,14 @@ Section proof.
iDestruct
"Hs"
as
(
xs
''
hd
'''
)
"[>Hhd [>Hxs HRs]]"
.
iDestruct
"HRs"
as
(
m
'
)
"[>Hom HRs]"
.
iDestruct
(
ev_map_witness
_
_
_
m
'
with
"[Hevm Hom]"
)
as
%?
;
first
by
iFrame
.
case
Hlk
'
:
(
m
'
!!
xhd
)
=>
[
xx
|
];
last
first
.
{
rewrite
Hlk
'
in
H
.
inversion
H
.
}
iDestruct
(
big_sepM_delete_later
(
perR
_
)
m
'
with
"HRs"
)
as
"[Hp HRm]"
=>
//.
iDestruct
"Hp"
as
(
v
''
)
"[>% [Hpinv' >Hhd'']]"
.
inversion
H1
.
subst
.
rewrite
Hlk
'
in
H
.
iDestruct
"Hp"
as
(
v
''
)
"[>% [Hpinv' >Hhd'']]"
.
inversion
H
.
subst
.
destruct
H3
as
[
_
H3
].
apply
(
inj
to_agree
)
in
H3
.
fold_leibniz
.
subst
.
iDestruct
"Hpinv'"
as
([[[[
γ
x
'
γ
1
'
]
γ
3
'
]
γ
4
'
]
γ
q
'
]
p
''''
)
"[>% [>#Hevm2 Hps]]"
.
inversion
H
2
.
subst
.
inversion
H
0
.
subst
.
destruct
(
decide
(
γ
1
=
γ
1
'
∧
γ
x
=
γ
x
'
∧
γ
3
=
γ
3
'
∧
γ
4
=
γ
4
'
∧
γ
q
=
γ
q
'
))
as
[[
?
[
?
[
?
[
?
?
]]]]
|
Hneq
];
subst
.
{
...
...
theories/peritem.v
View file @
535a2fc5
...
...
@@ -36,7 +36,7 @@ Section defs.
iApply
IHxs
'
=>
//.
Qed
.
Definition
perR
'
hd
v
v
'
:=
(
⌜
v
=
((
∅
:
unitR
),
DecA
gree
v
'
)
⌝
∗
R
v
'
∗
allocated
hd
)
%
I
.
Definition
perR
'
hd
v
v
'
:=
(
⌜
v
=
((
∅
:
unitR
),
to_a
gree
v
'
)
⌝
∗
R
v
'
∗
allocated
hd
)
%
I
.
Definition
perR
hd
v
:=
(
∃
v
'
,
perR
'
hd
v
v
'
)
%
I
.
Definition
allR
γ
:=
(
∃
m
:
evmapR
loc
val
unitR
,
own
γ
(
●
m
)
∗
[
∗
map
]
hd
↦
v
∈
m
,
perR
hd
v
)
%
I
.
...
...
@@ -164,7 +164,7 @@ Lemma new_stack_spec' Φ RI:
iDestruct
(
big_sepM_insert_later
(
perR
R
)
m
)
as
"H"
=>
//.
iSplitL
"Hox"
.
{
rewrite
/
evs
/
ev
.
eauto
.
}
iExists
(
<
[
hd
'
:=
((),
DecA
gree
x
)]
>
m
).
iExists
(
<
[
hd
'
:=
((),
to_a
gree
x
)]
>
m
).
iFrame
.
iApply
"H"
.
iFrame
.
iExists
x
.
iFrame
.
rewrite
/
allocated
.
iSplitR
"Hhd'1"
;
auto
.
}
...
...
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