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
Iris
examples
Commits
3bb9aedd
Commit
3bb9aedd
authored
Sep 22, 2016
by
Zhen Zhang
Browse files
rewrite: hope i got things right this time
parent
889d57f1
Changes
1
Hide whitespace changes
Inline
Side-by-side
srv.v
View file @
3bb9aedd
...
...
@@ -7,20 +7,20 @@ From iris.algebra Require Import upred frac excl dec_agree upred_big_op gset gma
From
iris
.
tests
Require
Import
atomic
treiber_stack
.
From
flatcomb
Require
Import
misc
.
Definition
doOp
:
val
:
=
λ
:
"f"
"p"
,
Definition
doOp
(
f
:
val
)
:
val
:
=
λ
:
"p"
,
match
:
!
"p"
with
InjL
"x"
=>
"p"
<-
InjR
(
"f"
"x"
)
InjL
"x"
=>
"p"
<-
InjR
(
f
"x"
)
|
InjR
"_"
=>
#()
end
.
Definition
loop
:
val
:
=
rec
:
"loop"
"p"
"f"
"s"
"lk"
:
=
Definition
loop
(
f
:
val
)
:
val
:
=
rec
:
"loop"
"p"
"s"
"lk"
:
=
match
:
!
"p"
with
InjL
"_"
=>
if
:
CAS
"lk"
#
false
#
true
then
iter
(
doOp
"f"
)
"s"
else
"loop"
"p"
"f"
"s"
"lk"
then
iter
(
doOp
f
)
"s"
else
"loop"
"p"
"s"
"lk"
|
InjR
"r"
=>
"r"
end
.
...
...
@@ -31,240 +31,127 @@ Definition install : val :=
push
"s"
"p"
;;
"p"
.
Definition
flat
:
val
:
=
Definition
flat
(
f
:
val
)
:
val
:
=
λ
:
"f"
,
let
:
"lk"
:
=
ref
(#
false
)
in
let
:
"s"
:
=
new_stack
#()
in
λ
:
"x"
,
let
:
"p"
:
=
install
"x"
"s"
in
loop
"p"
"
f
"
"s"
"lk"
.
loop
f
"
p
"
"s"
"lk"
.
Global
Opaque
doOp
install
loop
flat
.
Definition
srvR
:
=
prodR
fracR
(
dec_agreeR
val
).
Class
srvG
Σ
:
=
SrvG
{
srv_G
:
>
inG
Σ
srvR
}.
Definition
srv
Σ
:
gFunctors
:
=
#[
GFunctor
(
constRF
srvR
)].
Definition
tokR
:
=
evmapR
loc
(
gname
*
gname
*
gname
*
gname
).
Class
srvG
Σ
:
=
SrvG
{
srv_G
:
>
inG
Σ
srvR
;
tok_G
:
>
inG
Σ
(
authR
tokR
)
;
}.
Definition
srv
Σ
:
gFunctors
:
=
#[
GFunctor
(
constRF
srvR
)
;
GFunctor
(
constRF
(
authR
tokR
))
].
Instance
subG_srv
Σ
{
Σ
}
:
subG
srv
Σ
Σ
→
srvG
Σ
.
Proof
.
intros
[?%
subG_inG
_
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Proof
.
intros
[?%
subG_inG
[?%
subG_inG
_
]%
subG_inv
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Section
proof
.
Context
`
{!
heapG
Σ
,
!
lockG
Σ
,
!
evidenceG
loc
val
Σ
,
!
srvG
Σ
}
(
N
:
namespace
).
Definition
p_inv
(
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
(
γ
m
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
(
Q
:
val
→
val
→
Prop
)
(
v
:
val
)
:
iProp
Σ
:
=
((
∃
(
y
:
val
),
v
=
InjRV
y
★
own
γ
1
(
Excl
())
★
own
γ
3
(
Excl
()))
∨
(
∃
(
x
:
val
),
v
=
InjLV
x
★
own
γ
x
((
1
/
2
)%
Qp
,
DecAgree
x
)
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
()))
∨
(
∃
(
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
.
(
∃
(
p
:
loc
)
(
q
:
Qp
),
v
=
#
p
★
own
γ
m
(
◯
({[
p
:
=
(
q
,
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
))
]}
:
tokR
))
★
((
∃
(
y
:
val
),
p
↦
InjRV
y
★
own
γ
1
(
Excl
())
★
own
γ
3
(
Excl
()))
∨
(
∃
(
x
:
val
),
p
↦
InjLV
x
★
own
γ
x
((
1
/
2
)%
Qp
,
DecAgree
x
)
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
()))
∨
(
∃
(
x
:
val
),
p
↦
InjLV
x
★
own
γ
x
((
1
/
4
)%
Qp
,
DecAgree
x
)
★
own
γ
2
(
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
.
Definition
p_inv_R
γ
2
Q
v
:
iProp
Σ
:
=
(
∃
γ
x
γ
1
γ
3
γ
4
:
gname
,
p_inv
γ
x
γ
1
γ
2
γ
3
γ
4
Q
v
)%
I
.
Definition
p_inv_R
γ
m
γ
2
Q
v
:
iProp
Σ
:
=
(
∃
γ
x
γ
1
γ
3
γ
4
:
gname
,
p_inv
γ
m
γ
x
γ
1
γ
2
γ
3
γ
4
Q
v
)%
I
.
Definition
srv_inv
(
γ
γ
2
:
gname
)
(
s
:
loc
)
(
Q
:
val
→
val
→
Prop
)
:
iProp
Σ
:
=
(
∃
xs
,
is_stack'
(
p_inv_R
γ
2
Q
)
xs
s
γ
)%
I
.
Definition
srv_
stack_
inv
(
γ
γ
m
γ
2
:
gname
)
(
s
:
loc
)
(
Q
:
val
→
val
→
Prop
)
:
iProp
Σ
:
=
(
∃
xs
,
is_stack'
(
p_inv_R
γ
m
γ
2
Q
)
xs
s
γ
)%
I
.
Instance
p_inv_timeless
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
:
TimelessP
(
p_inv
γ
x
γ
1
γ
2
γ
3
γ
4
p
Q
).
Proof
.
apply
_
.
Qed
.
Definition
srv_m_inv
γ
m
:
=
(
∃
m
,
own
γ
m
(
●
m
))%
I
.
Instance
srv_inv_timeless
γ
γ
2
s
Q
:
TimelessP
(
srv_inv
γ
γ
2
s
Q
).
Lemma
install_push_spec
(
Φ
:
val
→
iProp
Σ
)
(
Q
:
val
→
val
→
Prop
)
(
p
:
loc
)
(
γ
γ
m
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
)
(
s
:
loc
)
(
x
:
val
)
:
heapN
⊥
N
→
heap_ctx
★
inv
N
(
srv_stack_inv
γ
γ
m
γ
2
s
Q
)
★
own
γ
x
((
1
/
2
)%
Qp
,
DecAgree
x
)
★
own
γ
m
(
◯
({[
p
:
=
((
1
/
2
)%
Qp
,
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
))
]}))
★
p
↦
InjLV
x
★
own
γ
1
(
Excl
())
★
own
γ
4
(
Excl
())
★
(
True
-
★
Φ
#())
⊢
WP
push
#
s
#
p
{{
Φ
}}.
Proof
.
apply
uPred
.
exist_timeless
.
move
=>
x
.
apply
is_stack'_timeless
.
move
=>
v
.
apply
_
.
iIntros
(
HN
)
"(#Hh & #? & Hx & Hpm & Hp & Ho1 & Ho2 & HΦ)"
.
rewrite
/
srv_stack_inv
.
iDestruct
(
push_spec
N
(
p_inv_R
γ
m
γ
2
Q
)
with
"[-HΦ]"
)
as
"Hpush"
=>//.
-
iFrame
"Hh"
.
iSplitL
"Hx Hp Hpm Ho1 Ho2"
;
last
eauto
.
iExists
γ
x
,
γ
1
,
γ
3
,
γ
4
.
iExists
p
,
(
1
/
2
)%
Qp
.
iSplit
=>//.
iFrame
"Hpm"
.
iRight
.
iLeft
.
iExists
x
.
iFrame
.
-
iApply
wp_wand_r
.
iSplitL
"Hpush"
;
first
done
.
iIntros
(?).
iIntros
"%"
.
inversion
H
.
by
iApply
"HΦ"
.
Qed
.
(* Lemma push_spec *)
(* (Φ: val → iProp Σ) (Q: val → val → Prop) *)
(* (p: loc) (γx γ1 γ2 γ3 γ4: gname) *)
(* (γhd γgn: gname) (s: loc) (x: val) : *)
(* heapN ⊥ N → *)
(* heap_ctx ★ inv N (srv_inv γhd γgn s Q) ★ own γx ((1/2)%Qp, DecAgree x) ★ *)
(* p ↦ InjLV x ★ own γ1 (Excl ()) ★ own γ4 (Excl ()) ★ (True -★ Φ #()) *)
(* ⊢ WP push #s #p {{ Φ }}. *)
(* Proof. *)
(* iIntros (HN) "(#Hh & #Hsrv & Hp & Hx & Ho1 & Ho4 & HΦ)". *)
(* iDestruct (push_atomic_spec N s #p with "Hh") as "Hpush"=>//. *)
(* rewrite /push_triple /atomic_triple. *)
(* iSpecialize ("Hpush" $! (p ↦ InjLV x ★ own γ1 (Excl ()) ★ own γ4 (Excl ()) ★ *)
(* own γx ((1/2)%Qp, DecAgree x))%I *)
(* (fun _ ret => ret = #())%I with "[]"). *)
(* - iIntros "!#". iIntros "(Hp & Hx & Ho1 & Ho4)". *)
(* (* open the invariant *) *)
(* iInv N as (hds gnm) ">(Hohd & Hogn & Hxs & Hhd & Hps)" "Hclose". *)
(* iDestruct "Hxs" as (xs) "(Hs & Hgn)". *)
(* (* mask magic *) *)
(* iApply pvs_intro'. *)
(* { apply ndisj_subseteq_difference; auto. } *)
(* iIntros "Hvs". *)
(* iExists (map (λ x : loc, #x) xs). *)
(* iFrame "Hs". iSplit. *)
(* + (* provide a way to rollback *) *)
(* iIntros "Hl'". *)
(* iVs "Hvs". iVs ("Hclose" with "[-Hp Hx Ho1 Ho4]"); last by iFrame. *)
(* iNext. rewrite /srv_inv. iExists hds, gnm. *)
(* iFrame. iExists xs. by iFrame. *)
(* + (* provide a way to commit *) *)
(* iIntros (?) "[% Hs]". subst. *)
(* iVs "Hvs". iVs ("Hclose" with "[-]"); last done. *)
(* iNext. rewrite /srv_inv. iExists hds, (gnm ∪ {[ p := DecAgree (γx, γ1, γ2, γ3, γ4) ]}). *)
(* iFrame. *)
(* iClear "Hogn". *)
(* iAssert (own γgn (● (gnm ∪ {[p := DecAgree (γx, γ1, γ2, γ3, γ4)]})) ★ *)
(* own γgn (◯ {[ p := DecAgree (γx, γ1, γ2, γ3, γ4) ]}))%I as "[Hogn' Hfrag]". *)
(* { admit. } *)
(* iFrame. iSplitL "Hs Hgn". *)
(* { iExists (p::xs). *)
(* iFrame. admit. } *)
(* iSplitL "Hhd". *)
(* { admit. } *)
(* iAssert (p_inv' (DecAgree (γx, γ1, γ2, γ3, γ4)) p Q) with "[Hp Hx Ho1 Ho4]" as "Hinvp". *)
(* { rewrite /p_inv' /p_inv. iRight. iLeft. iExists x. by iFrame. } *)
(* admit. *)
(* - iApply wp_wand_r. iSplitR "HΦ". *)
(* + iApply "Hpush". by iFrame. *)
(* + iIntros (?) "H". iDestruct "H" as (?) "%". subst. by iApply "HΦ". *)
(* Admitted. *)
(* Lemma install_spec *)
(* (Φ: val → iProp Σ) (Q: val → val → Prop) *)
(* (x: val) (γhd γgn: gname) (s: loc): *)
(* heapN ⊥ N → *)
(* heap_ctx ★ inv N (srv_inv γhd γgn s Q) ★ *)
(* (∀ (p: loc) (γx γ1 γ2 γ3 γ4: gname), *)
(* own γ2 (Excl ()) -★ own γ3 (Excl ()) -★ own γgn (◯ {[ p := DecAgree (γx, γ1, γ2, γ3, γ4) ]}) -★ *)
(* own γx ((1/2)%Qp, DecAgree x) -★ Φ #p) *)
(* ⊢ WP install x #s {{ Φ }}. *)
(* Proof. *)
(* iIntros (HN) "(#Hh & #Hsrv & HΦ)". *)
(* wp_seq. wp_let. wp_alloc p as "Hl". *)
(* iVs (own_alloc (Excl ())) as (γ1) "Ho1"; first done. *)
(* iVs (own_alloc (Excl ())) as (γ2) "Ho2"; first done. *)
(* iVs (own_alloc (Excl ())) as (γ3) "Ho3"; first done. *)
(* iVs (own_alloc (Excl ())) as (γ4) "Ho4"; first done. *)
(* iVs (own_alloc (1%Qp, DecAgree x)) as (γx) "Hx"; first done. *)
(* iDestruct (own_update with "Hx") as "==>[Hx1 Hx2]". *)
(* { by apply pair_l_frac_op_1'. } *)
(* wp_let. wp_bind ((push _) _). *)
(* iApply push_spec=>//. *)
(* iFrame "Hh Hsrv Hx1 Hl Ho1 Ho4". *)
(* iIntros "_". wp_seq. iVsIntro. *)
(* iSpecialize ("HΦ" $! p γx γ1 γ2 γ3 γ4). *)
(* iAssert (own γgn (◯ {[p := DecAgree (γx, γ1, γ2, γ3, γ4)]})) as "Hfrag". *)
(* { admit. } *)
(* iApply ("HΦ" with "Ho2 Ho3 Hfrag Hx2"). *)
(* Admitted. *)
(* Definition pinv_sub RI γx γ1 γ2 γ3 γ4 p Q := (RI ⊣⊢ ∃ Rf, Rf ★ p_inv γx γ1 γ2 γ3 γ4 p Q)%I. *)
(* Lemma doOp_spec Φ (f: val) (RI: iProp Σ) γx γ1 γ2 γ3 γ4 p Q `{TimelessP _ RI}: *)
(* heapN ⊥ N → pinv_sub RI γx γ1 γ2 γ3 γ4 p Q → *)
(* heap_ctx ★ inv N RI ★ own γ2 (Excl ()) ★ *)
(* □ (∀ x:val, WP f x {{ v, ■ Q x v }})%I ★ (own γ2 (Excl ()) -★ Φ #()) *)
(* ⊢ WP doOp f #p {{ Φ }}. *)
(* Proof. *)
(* iIntros (HN Hsub) "(#Hh & #HRI & Ho2 & #Hf & HΦ)". *)
(* wp_seq. wp_let. wp_bind (! _)%E. *)
(* iInv N as ">H" "Hclose". *)
(* iDestruct (Hsub with "H") as (Rf) "[HRf [Hp | [Hp | [Hp | Hp]]]]". *)
(* - iDestruct "Hp" as (y) "(Hp & Ho1 & Ho3)". *)
(* wp_load. iVs ("Hclose" with "[HRf Hp Ho1 Ho3]"). *)
(* { iNext. iApply Hsub. iExists Rf. iFrame "HRf". *)
(* iLeft. iExists y. by iFrame. } *)
(* iVsIntro. wp_match. by iApply "HΦ". *)
(* - iDestruct "Hp" as (x) "(Hp & Hx & Ho1 & Ho4)". *)
(* wp_load. *)
(* iAssert (|=r=> own γx (((1 / 4)%Qp, DecAgree x) ⋅ ((1 / 4)%Qp, DecAgree x)))%I with "[Hx]" as "==>[Hx1 Hx2]". *)
(* { iDestruct (own_update with "Hx") as "Hx"; last by iAssumption. *)
(* replace ((1 / 2)%Qp) with (1/4 + 1/4)%Qp; last by apply Qp_div_S. *)
(* by apply pair_l_frac_op'. } *)
(* iVs ("Hclose" with "[HRf Hp Hx1 Ho2 Ho4]"). *)
(* { iNext. iApply Hsub. iExists Rf. iFrame "HRf". *)
(* iRight. iRight. iLeft. iExists x. by iFrame. } *)
(* iVsIntro. wp_match. *)
(* wp_bind (f _). iApply wp_wand_r. *)
(* iSplitR; first by iApply "Hf". *)
(* iIntros (y) "%". *)
(* iInv N as ">H" "Hclose". *)
(* iDestruct (Hsub with "H") as (Rf') "[HRf [Hp | [Hp | [Hp | Hp]]]]". *)
(* + admit. *)
(* + admit. *)
(* + iDestruct "Hp" as (x') "(Hp & Hx & Ho2 & Ho4)". *)
(* destruct (decide (x = x')) as [->|Hneq]; last by admit. *)
(* iCombine "Hx2" "Hx" as "Hx". *)
(* iDestruct (own_update with "Hx") as "==>Hx"; first by apply pair_l_frac_op. *)
(* rewrite Qp_div_S. *)
(* wp_store. iVs ("Hclose" with "[HRf Hp Hx Ho1 Ho4]"). *)
(* { iNext. iApply Hsub. iExists Rf'. iFrame "HRf". *)
(* iRight. iRight. iRight. iExists x', y. *)
(* by iFrame. } *)
(* iVsIntro. by iApply "HΦ". *)
(* + admit. *)
(* - admit. *)
(* - admit. *)
(* Admitted. *)
Lemma
install_spec
(
Φ
:
val
→
iProp
Σ
)
(
Q
:
val
→
val
→
Prop
)
(
x
:
val
)
(
γ
2
γ
m
γ
:
gname
)
(
s
:
loc
)
:
heapN
⊥
N
→
heap_ctx
★
inv
N
(
srv_stack_inv
γ
γ
m
γ
2
s
Q
)
★
inv
N
(
srv_m_inv
γ
m
)
★
(
∀
(
p
:
loc
)
(
γ
x
γ
1
γ
2
γ
3
γ
4
:
gname
),
own
γ
3
(
Excl
())
-
★
own
γ
m
(
◯
{[
p
:
=
((
1
/
2
)%
Qp
,
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
))
]})
-
★
own
γ
x
((
1
/
2
)%
Qp
,
DecAgree
x
)
-
★
Φ
#
p
)
⊢
WP
install
x
#
s
{{
Φ
}}.
Proof
.
iIntros
(
HN
)
"(#Hh & #? & #? & HΦ)"
.
wp_seq
.
wp_let
.
wp_alloc
p
as
"Hl"
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
1
)
"Ho1"
;
first
done
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
3
)
"Ho3"
;
first
done
.
iVs
(
own_alloc
(
Excl
()))
as
(
γ
4
)
"Ho4"
;
first
done
.
iVs
(
own_alloc
(
1
%
Qp
,
DecAgree
x
))
as
(
γ
x
)
"Hx"
;
first
done
.
iDestruct
(
own_update
with
"Hx"
)
as
"==>[Hx1 Hx2]"
;
first
by
apply
pair_l_frac_op_1'
.
wp_let
.
wp_bind
((
push
_
)
_
).
iApply
install_push_spec
=>//.
iFrame
"#"
.
iAssert
(
own
γ
m
(
◯
{[
p
:
=
((
1
/
2
)%
Qp
,
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
))]}
⋅
◯
{[
p
:
=
((
1
/
2
)%
Qp
,
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
))]}))
as
"[Hfrag1 Hfrag2]"
.
{
admit
.
}
iFrame
"Hx1 Hfrag1 Hl Ho1 Ho4"
.
iIntros
"_"
.
wp_seq
.
iVsIntro
.
iSpecialize
(
"HΦ"
$!
p
γ
x
γ
1
γ
2
γ
3
γ
4
with
"[-Hx2 Hfrag2]"
)=>//.
iApply
(
"HΦ"
with
"Hfrag2 Hx2"
).
Admitted
.
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
(
γ
γ
m
γ
2
:
gname
)
xs
:
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
())
★
is_list
hd
xs
★
own
γ
hd
(
◯
({[
hd
]}
:
hdsetR
))
★
(
own
γ
2
(
Excl
())
-
★
Φ
#())
⊢
WP
doOp
f
{{
f'
,
WP
iter'
#
hd
f'
{{
Φ
}}
}}.
heap_ctx
★
inv
N
(
srv_
stack_inv
γ
γ
m
γ
2
s
Q
)
★
□
(
∀
x
:
val
,
WP
f
x
{{
v
,
■
Q
x
v
}})%
I
★
own
γ
2
(
Excl
())
★
is_list
'
γ
hd
xs
★
(
own
γ
2
(
Excl
())
-
★
Φ
#())
⊢
WP
iter'
#
hd
(
doOp
f
)
{{
Φ
}}.
Proof
.
iIntros
(
HN
)
"(#Hh & #? & #Hf & Hxs & Hhd & Ho2 & HΦ)"
.
rewrite
/
doOp
.
wp_let
.
iL
ö
b
as
"IH"
.
wp_rec
.
wp_let
.
wp_bind
(!
_
)%
E
.
destruct
xs
as
[|
x
xs'
].
-
simpl
.
iDestruct
"Hhd"
as
(
q
)
"Hhd"
.
wp_load
.
wp_match
.
by
iApply
"HΦ"
.
-
simpl
.
iDestruct
"Hhd"
as
(
hd'
q
)
"[[Hhd1 Hhd2] Hhd']"
.
wp_load
.
wp_match
.
wp_proj
.
wp_let
.
wp_bind
(!
_
)%
E
.
iInv
N
as
(
hds
gnm
)
">(Hohd & Hogn & Hxse & Hhds & Hps)"
"Hclose"
.
iAssert
(
∃
(
p
:
loc
)
γ
s
,
x
=
#
p
★
p_inv'
γ
2
γ
s
p
Q
)%
I
as
"Hx"
.
{
admit
.
}
iDestruct
"Hx"
as
(
p
γ
s
)
"[% Hp]"
.
subst
.
rewrite
/
p_inv'
.
destruct
γ
s
as
[[[[
γ
x
γ
1
]
γ
3
]
γ
4
]|].
iDestruct
"Hp"
as
"[Hp | [Hp | [ Hp | Hp]]]"
=>//.
+
admit
.
+
iDestruct
"Hp"
as
(
x
)
"(Hp & Hx & Ho1 & Ho4)"
.
iAssert
(|=
r
=>
own
γ
x
(((
1
/
4
)%
Qp
,
DecAgree
x
)
⋅
((
1
/
4
)%
Qp
,
DecAgree
x
)))%
I
with
"[Hx]"
as
"==>[Hx1 Hx2]"
.
{
iDestruct
(
own_update
with
"Hx"
)
as
"Hx"
;
last
by
iAssumption
.
replace
((
1
/
2
)%
Qp
)
with
(
1
/
4
+
1
/
4
)%
Qp
;
last
by
apply
Qp_div_S
.
by
apply
pair_l_frac_op'
.
}
wp_load
.
iVs
(
"Hclose"
with
"[-Ho1 Hx2 Hhd2 HΦ]"
).
{
iNext
.
iExists
hds
,
gnm
.
by
iFrame
.
}
iVsIntro
.
wp_match
.
wp_bind
(
f
_
).
iApply
wp_wand_r
.
iSplitR
;
first
by
iApply
"Hf"
.
iIntros
(
y
)
"Q"
.
iInv
N
as
(
hds'
gnm'
)
">(Hohd & Hogn & Hxse & Hhds & Hps)"
"Hclose"
.
iAssert
(
p_inv'
γ
2
(
DecAgree
(
γ
x
,
γ
1
,
γ
3
,
γ
4
))
p
Q
)%
I
as
"Hp"
.
{
admit
.
}
rewrite
/
p_inv'
.
iDestruct
"Hp"
as
"[Hp | [Hp | [ Hp | Hp]]]"
.
*
admit
.
*
admit
.
*
iDestruct
"Hp"
as
(
x'
)
"(Hp & Hx' & Ho2 & Ho4)"
.
destruct
(
decide
(
x
=
x'
))
as
[->|
Hneq
]
;
last
by
admit
.
iCombine
"Hx2"
"Hx'"
as
"Hx"
.
iDestruct
(
own_update
with
"Hx"
)
as
"==>Hx"
;
first
by
apply
pair_l_frac_op
.
rewrite
Qp_div_S
.
wp_store
.
iVs
(
"Hclose"
with
"[-Ho2 HΦ Hhd2]"
).
{
iNext
.
iExists
hds'
,
gnm'
.
by
iFrame
.
}
iVsIntro
.
wp_seq
.
wp_proj
.
rewrite
/
doOp
.
iAssert
(
is_list
hd'
xs'
)
as
"Hl"
.
{
admit
.
}
iAssert
(
∃
(
hd'0
:
loc
)
(
q0
:
Qp
),
hd
↦
{
q0
}
SOMEV
(#
p
,
#
hd'0
)
★
is_list
hd'0
xs'
)%
I
with
"[Hhd2 Hl]"
as
"He"
.
{
iExists
hd'
,
(
q
/
2
)%
Qp
.
by
iFrame
.
}
iAssert
(
own
γ
hd
(
◯
{[
hd
]}))
as
"Hfrag"
.
{
admit
.
}
iSpecialize
(
"IH"
with
"Ho2 He Hfrag HΦ"
).
admit
.
*
admit
.
+
admit
.
+
admit
.
+
by
iExFalso
.
iIntros
(
HN
)
"(#Hh & #? & #Hf & Ho2 & Hlist' & HΦ)"
.
rewrite
/
doOp
.
iApply
pvs_wp
.
iDestruct
(
dup_is_list'
with
"[Hlist']"
)
as
"==>[Hl1 Hl2]"
;
first
by
iFrame
.
iDestruct
(
dup_is_list'
with
"[Hl2]"
)
as
"==>[Hl2 Hl3]"
;
first
by
iFrame
.
iDestruct
(
iter'_spec
_
(
p_inv_R
γ
m
γ
2
Q
)
_
γ
s
(
is_list'
γ
hd
xs
★
own
γ
2
(
Excl
()))%
I
xs
hd
(
λ
:
"p"
,
match
:
!
"p"
with
InjL
"x"
=>
"p"
<-
SOME
(
f
"x"
)
|
InjR
"_"
=>
#()
end
)%
V
(
λ
:
"p"
,
match
:
!
"p"
with
InjL
"x"
=>
"p"
<-
SOME
(
f
"x"
)
|
InjR
"_"
=>
#()
end
)
with
"[-Hl1]"
)
as
"Hiter'"
=>//.
-
rewrite
/
f_spec
.
iIntros
(
Φ
'
x
_
Hin
)
"(#Hh & #? & (Hls & Ho2) & HΦ')"
.
wp_let
.
wp_bind
(!
_
)%
E
.
iInv
N
as
(
xs'
)
">Hs"
"Hclose"
.
iDestruct
"Hs"
as
(
hd'
)
"[Hhd [Hxs HRs]]"
.
(* now I know x conforms to p_inv_R *)
admit
.
-
apply
to_of_val
.
-
iFrame
"#"
.
iFrame
"Hl2 Hl3 Ho2"
.
iIntros
"[_ H]"
.
by
iApply
"HΦ"
.
-
done
.
Admitted
.
Lemma
loop_iter_spec
Φ
(
f
:
val
)
(
s
:
loc
)
Q
(
γ
hd
γ
gn
γ
2
:
gname
)
:
...
...
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