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
Jonas Kastberg
actris
Commits
1b9fc4de
Commit
1b9fc4de
authored
May 03, 2019
by
Jonas Kastberg Hinrichsen
Browse files
WIP List sort example
parent
4ce5eb98
Changes
3
Hide whitespace changes
Inline
Side-by-side
theories/encodings/list_enc.v
View file @
1b9fc4de
...
...
@@ -3,10 +3,18 @@ From iris.heap_lang Require Import assert.
From
osiris
Require
Export
encodable
.
(** Immutable ML-style functional lists *)
Fixpoint
is_list_enc
`
{
EncDec
T
}
(
hd
:
val
)
(
xs
:
list
T
)
:
Prop
:
=
Instance
pair_encodable
`
{
Encodable
A
,
Encodable
B
}
:
Encodable
(
A
*
B
)
:
=
λ
p
,
(
encode
p
.
1
,
encode
p
.
2
)%
V
.
Instance
option_encodable
`
{
Encodable
A
}
:
Encodable
(
option
A
)
:
=
λ
mx
,
match
mx
with
Some
x
=>
SOMEV
(
encode
x
)
|
None
=>
NONEV
end
%
V
.
Instance
list_encodable
`
{
Encodable
A
}
:
Encodable
(
list
A
)
:
=
fix
go
xs
:
=
let
_
:
Encodable
_
:
=
@
go
in
match
xs
with
|
[]
=>
hd
=
NONEV
|
x
::
xs
=>
∃
hd'
,
hd
=
SOMEV
(
encode
x
,
hd'
)
∧
is_list_enc
hd'
xs
|
[]
=>
encode
None
|
x
::
xs
=>
encode
(
Some
(
x
,
xs
))
end
.
Definition
lnil
:
val
:
=
λ
:
<>,
NONEV
.
...
...
@@ -67,135 +75,109 @@ Section lists.
Context
`
{
heapG
Σ
}.
Implicit
Types
i
:
nat
.
Context
`
{
EncDec
T
}.
Implicit
Types
v
:
val
.
Implicit
Types
x
:
T
.
Implicit
Types
xs
:
list
T
.
Lemma
lnil_spec
:
{{{
True
}}}
lnil
#()
{{{
hd
,
RET
hd
;
⌜
is_list_enc
hd
[]
⌝
}}}.
{{{
True
}}}
lnil
#()
{{{
RET
encode
[]
;
True
}}}.
Proof
.
iIntros
(
Φ
_
)
"HΦ"
.
wp_lam
.
by
iApply
"HΦ"
.
Qed
.
Lemma
lcons_spec
hd
xs
x
:
{{{
⌜
is_list_enc
hd
xs
⌝
}}}
lcons
(
encode
x
)
hd
{{{
hd'
,
RET
hd'
;
⌜
is_list_enc
hd'
(
x
::
xs
)
⌝
}}}.
Proof
.
iIntros
(
Φ
?)
"HΦ"
.
wp_lam
.
wp_pures
.
iApply
"HΦ"
.
simpl
;
eauto
10
.
Qed
.
Lemma
lcons_spec
x
xs
:
{{{
True
}}}
lcons
(
encode
x
)
(
encode
xs
)
{{{
RET
(
encode
(
x
::
xs
))
;
True
}}}.
Proof
.
iIntros
(
Φ
_
)
"HΦ"
.
wp_lam
.
wp_pures
.
by
iApply
"HΦ"
.
Qed
.
Lemma
lhead_spec
hd
x
xs
:
{{{
⌜
is_list_enc
hd
(
x
::
xs
)
⌝
}}}
lhead
hd
{{{
RET
encode
x
;
True
}}}.
Proof
.
iIntros
(
Φ
(
hd'
&->&?))
"HΦ"
.
wp_lam
.
wp_pures
.
by
iApply
"HΦ"
.
Qed
.
Lemma
lhead_spec
x
xs
:
{{{
True
}}}
lhead
(
encode
(
x
::
xs
))
{{{
RET
encode
x
;
True
}}}.
Proof
.
iIntros
(
Φ
_
)
"HΦ"
.
wp_lam
.
wp_pures
.
by
iApply
"HΦ"
.
Qed
.
Lemma
ltail_spec
hd
xs
x
:
{{{
⌜
is_list_enc
hd
(
x
::
xs
)
⌝
}}}
ltail
hd
{{{
hd'
,
RET
hd'
;
⌜
is_list_enc
hd'
xs
⌝
}}}.
Proof
.
iIntros
(
Φ
(
hd'
&->&?))
"HΦ"
.
wp_lam
.
wp_pures
.
by
iApply
"HΦ"
.
Qed
.
Lemma
ltail_spec
x
xs
:
{{{
True
}}}
ltail
(
encode
(
x
::
xs
))
{{{
RET
encode
xs
;
True
}}}.
Proof
.
iIntros
(
Φ
_
)
"HΦ"
.
wp_lam
.
wp_pures
.
by
iApply
"HΦ"
.
Qed
.
Lemma
llookup_spec
i
hd
xs
x
:
Lemma
llookup_spec
i
xs
x
:
xs
!!
i
=
Some
x
→
{{{
⌜
is_list_enc
hd
xs
⌝
}}}
llookup
#
i
hd
{{{
True
}}}
llookup
#
i
(
encode
xs
)
{{{
RET
encode
x
;
True
}}}.
Proof
.
iIntros
(
Hi
Φ
Hl
)
"HΦ"
.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
hd
i
Hi
Hl
)
;
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
i
Hi
Hl
)
;
destruct
i
as
[|
i
]=>
//
;
simplify_eq
/=.
{
wp_lam
.
wp_pures
.
by
iApply
(
lhead_spec
with
"[//]"
).
}
wp_lam
.
wp_pures
.
wp_apply
(
ltail_spec
with
"[//]"
)
;
iIntros
(
hd'
?).
wp_let
.
wp_apply
(
ltail_spec
with
"[//]"
)
;
iIntros
(?).
wp_let
.
wp_op
.
rewrite
Nat2Z
.
inj_succ
-
Z
.
add_1_l
Z
.
add_simpl_l
.
by
iApply
"IH"
.
Qed
.
Lemma
linsert_spec
i
hd
xs
x
:
Lemma
linsert_spec
i
xs
x
:
is_Some
(
xs
!!
i
)
→
{{{
⌜
is_list_enc
hd
xs
⌝
}}}
linsert
#
i
(
encode
x
)
hd
{{{
hd'
,
RET
hd'
;
⌜
is_list_enc
hd'
(<[
i
:
=
x
]>
xs
)
⌝
}}}.
{{{
True
}}}
linsert
#
i
(
encode
x
)
(
encode
xs
)
{{{
RET
encode
(<[
i
:
=
x
]>
xs
)
;
True
}}}.
Proof
.
iIntros
([
w
Hi
]
Φ
Hl
)
"HΦ"
.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
hd
i
Hi
Hl
Φ
)
;
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
i
Hi
Hl
Φ
)
;
destruct
i
as
[|
i
]=>
//
;
simplify_eq
/=.
{
wp_lam
.
wp_pures
.
wp_apply
(
ltail_spec
with
"[//]"
).
iIntros
(
hd'
?).
{
wp_lam
.
wp_pures
.
wp_apply
(
ltail_spec
with
"[//]"
).
iIntros
(?).
wp_let
.
by
iApply
(
lcons_spec
with
"[//]"
).
}
wp_lam
;
wp_pures
.
wp_apply
(
ltail_spec
with
"[//]"
)
;
iIntros
(
hd'
?).
wp_let
.
wp_apply
(
ltail_spec
with
"[//]"
)
;
iIntros
(?).
wp_let
.
wp_op
.
rewrite
Nat2Z
.
inj_succ
-
Z
.
add_1_l
Z
.
add_simpl_l
.
wp_apply
(
"IH"
with
"[//] [//]"
).
iIntros
(
hd''
?).
wp_let
.
wp_apply
(
"IH"
with
"[//] [//]"
).
iIntros
(?).
wp_let
.
wp_apply
(
lhead_spec
with
"[//]"
)
;
iIntros
"_"
.
by
wp_apply
(
lcons_spec
with
"[//]"
).
Qed
.
Lemma
llist_member_spec
hd
xs
x
:
{{{
⌜
is_list_enc
hd
xs
⌝
}}}
llist_member
(
encode
x
)
hd
{{{
RET
#(
bool_decide
(
x
∈
xs
))
;
True
}}}.
Proof
.
iIntros
(
Φ
Hl
)
"HΦ"
.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
hd
Hl
)
;
simplify_eq
/=.
{
wp_lam
;
wp_pures
.
by
iApply
"HΦ"
.
}
destruct
Hl
as
(
hd'
&->&?).
wp_lam
;
wp_pures
.
destruct
(
bool_decide_reflect
(
encode
x'
=
encode
x
))
as
[
Heq
|?]
;
wp_if
.
{
rewrite
-
Heq
.
inversion
Heq
.
rewrite
(
bool_decide_true
(
_
∈
_
::
_
))
;
last
by
left
.
by
iApply
"HΦ"
.
}
wp_proj
.
wp_let
.
iApply
(
"IH"
with
"[//]"
).
destruct
(
bool_decide_reflect
(
x
∈
xs
)).
-
by
rewrite
bool_decide_true
;
last
by
right
.
-
by
rewrite
bool_decide_false
?elem_of_cons
;
last
naive_solver
.
Qed
.
(*
Lemma llist_member_spec
`{EqDecision T} `{Inj T} (xs : list T) (x : T) : *)
(*
{{{
True
}}}
*)
(*
llist_member (encode x)
(encode xs) *)
(*
{{{ RET #(bool_decide (x ∈ xs)); True }}}.
*)
(*
Proof.
*)
(*
iIntros (Φ Hl) "HΦ".
*)
(*
iInduction xs as [|x' xs] "IH" forall (Hl); simplify_eq/=.
*)
(*
{ wp_lam; wp_pures. by iApply "HΦ". }
*)
(*
wp_lam; wp_pures.
*)
(*
destruct (bool_decide_reflect (encode x' = encode x)) as [Heq|?]; wp_if.
*)
(*
{ rewrite -Heq. inversion Heq. rewrite (bool_decide_true (_ ∈ _ :: _))
. by iApply "HΦ". left.
by left. by iApply "HΦ". }
*)
(*
wp_proj. wp_let. iApply ("IH" with "[//]").
*)
(*
destruct (bool_decide_reflect (x ∈ xs)).
*)
(*
- by rewrite bool_decide_true; last by right.
*)
(*
- by rewrite bool_decide_false ?elem_of_cons; last naive_solver.
*)
(*
Qed.
*)
Lemma
lreplicate_spec
i
x
:
{{{
True
}}}
lreplicate
#
i
(
encode
x
)
{{{
hd
,
RET
hd
;
⌜
is_list_enc
hd
(
replicate
i
x
)
⌝
}}}.
{{{
RET
encode
(
replicate
i
x
)
;
True
}}}.
Proof
.
iIntros
(
Φ
_
)
"HΦ"
.
iInduction
i
as
[|
i
]
"IH"
forall
(
Φ
)
;
simpl
.
{
wp_lam
;
wp_pures
.
iApply
(
lnil_spec
with
"[//]"
).
iApply
"HΦ"
.
}
wp_lam
;
wp_pures
.
rewrite
Nat2Z
.
inj_succ
-
Z
.
add_1_l
Z
.
add_simpl_l
.
wp_apply
"IH"
;
iIntros
(
hd'
Hl
).
wp_let
.
by
iApply
(
lcons_spec
with
"[//]"
).
wp_apply
"IH"
;
iIntros
(
Hl
).
wp_let
.
by
iApply
(
lcons_spec
with
"[//]"
).
Qed
.
Lemma
llength_spec
hd
xs
:
{{{
⌜
is_list_enc
hd
xs
⌝
}}}
llength
hd
{{{
RET
#(
length
xs
)
;
True
}}}.
Lemma
llength_spec
xs
:
{{{
True
}}}
llength
(
encode
xs
)
{{{
RET
#(
length
xs
)
;
True
}}}.
Proof
.
iIntros
(
Φ
Hl
)
"HΦ"
.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
hd
Hl
Φ
)
;
simplify_eq
/=.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
Hl
Φ
)
;
simplify_eq
/=.
{
wp_lam
.
wp_match
.
by
iApply
"HΦ"
.
}
destruct
Hl
as
(
hd'
&->&?).
wp_lam
.
wp_match
.
wp_proj
.
wp_apply
(
"IH"
$!
hd'
with
"[//]"
)
;
iIntros
"_ /="
.
wp_op
.
wp_lam
.
wp_match
.
wp_proj
.
wp_apply
(
"IH"
with
"[//]"
)
;
iIntros
"_ /="
.
wp_op
.
rewrite
(
Nat2Z
.
inj_add
1
).
by
iApply
"HΦ"
.
Qed
.
Lemma
lsnoc_spec
hd
xs
x
:
{{{
⌜
is_list_enc
hd
xs
⌝
}}}
lsnoc
hd
(
encode
x
)
{{{
hd'
,
RET
hd'
;
⌜
is_list_enc
hd'
(
xs
++[
x
])
⌝
}}}.
Lemma
lsnoc_spec
xs
x
:
{{{
True
}}}
lsnoc
(
encode
xs
)
(
encode
x
)
{{{
RET
(
encode
(
xs
++[
x
]))
;
True
}}}.
Proof
.
iIntros
(
Φ
Hvs
)
"HΦ"
.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
hd
Hvs
Φ
).
-
simplify_eq
/=.
wp_rec
.
wp_pures
.
wp_lam
.
wp_pures
.
iApply
"HΦ"
.
simpl
.
eauto
10
.
-
destruct
Hvs
as
[
vs'
[->
Hvs'
]].
wp_rec
.
wp_pures
.
wp_bind
(
lsnoc
vs'
(
encode
x
)).
iApply
"IH"
.
eauto
.
iNext
.
iIntros
(
hd'
Hhd'
).
wp_pures
.
iApply
lcons_spec
.
eauto
.
iApply
"HΦ"
.
iInduction
xs
as
[|
x'
xs
]
"IH"
forall
(
Hvs
Φ
).
{
wp_rec
.
wp_pures
.
wp_lam
.
wp_pures
.
by
iApply
"HΦ"
.
}
wp_rec
.
wp_apply
"IH"
=>
//.
iIntros
(
_
).
by
wp_apply
lcons_spec
=>
//.
Qed
.
End
lists
.
\ No newline at end of file
theories/examples/list_sort_example.v
View file @
1b9fc4de
...
...
@@ -17,23 +17,23 @@ Section ListSortExample.
Definition
sorted
(
xs
:
list
Z
)
:
Prop
:
=
ordered
(
≤
)
xs
.
Definition
sorted_
hd
(
hd
:
val
)
(
xs
:
list
Z
)
:
Prop
:
=
is_list_enc
hd
xs
∧
sorted
xs
.
(*
Definition sorted_val (xs :
val
) : Prop :=
*)
(* sorted (encode xs). *)
Definition
sorted_of
(
xs
ys
:
list
Z
)
:
Prop
:
=
sorted
xs
∧
permutation
xs
ys
.
Definition
sorted_of_hd
(
hd
:
val
)
(
xs
ys
:
list
Z
)
:
Prop
:
=
is_list_enc
hd
xs
∧
sorted_of
xs
ys
.
(*
Definition sorted_of_hd (hd : val) (xs ys : list Z) : Prop :=
*)
(*
is_list_enc hd xs ∧ sorted_of xs ys.
*)
Definition
is_list_enc_ref
(
l
:
val
)
(
xs
:
list
Z
)
:
iProp
Σ
:
=
(
∃
l'
:
loc
,
∃
hd
:
val
,
⌜
l
=
#
l'
⌝
∧
l'
↦
hd
∗
⌜
is_list_enc
hd
xs
⌝
)%
I
.
(
∃
l'
:
loc
,
⌜
l
=
#
l'
⌝
∧
l'
↦
encode
xs
)%
I
.
Definition
sorted_ref
(
hd
:
val
)
(
xs
:
list
Z
)
:
iProp
Σ
:
=
(
is_list_enc_ref
hd
xs
∗
⌜
sorted
xs
⌝
)%
I
.
Definition
sorted_ref
(
lxs
:
val
)
(
xs
:
list
Z
)
:
iProp
Σ
:
=
(
is_list_enc_ref
lxs
xs
∗
⌜
sorted
xs
⌝
)%
I
.
Definition
sorted_of_ref
(
hd
:
val
)
(
xs
ys
:
list
Z
)
:
iProp
Σ
:
=
(
is_list_enc_ref
hd
xs
∗
⌜
sorted_of
xs
ys
⌝
)%
I
.
Definition
sorted_of_ref
(
lxs
:
val
)
(
xs
ys
:
list
Z
)
:
iProp
Σ
:
=
(
is_list_enc_ref
lxs
xs
∗
⌜
sorted_of
xs
ys
⌝
)%
I
.
Definition
lsplit
:
val
:
=
λ
:
"xs"
,
...
...
@@ -42,31 +42,28 @@ Section ListSortExample.
let
:
"zs"
:
=
ltail
!
"xs"
in
(
ref
"ys"
,
ref
"zs"
).
Lemma
lsplit_spec
(
hdx
:
val
)
(
xs
:
list
Z
)
:
{{{
is_list_enc_ref
hdx
xs
∗
(
∃
x
xs'
,
⌜
xs
=
x
::
xs
'
⌝
)}}}
lsplit
hdx
{{{
hdy
hdz
ys
zs
,
RET
(
hdy
,
hdz
)
;
is_list_enc_ref
hdx
xs
∗
is_list_enc_ref
hdy
ys
∗
is_list_enc_ref
hdz
zs
∗
⌜
xs
=
ys
++
zs
⌝
}}}.
Lemma
lsplit_spec
(
lxs
:
val
)
(
x
:
Z
)
(
xs
:
list
Z
)
:
{{{
is_list_enc_ref
lxs
(
x
::
xs
)
}}}
lsplit
lxs
{{{
lys
lzs
ys
zs
,
RET
(
lys
,
lzs
)
;
is_list_enc_ref
lxs
(
x
::
xs
)
∗
is_list_enc_ref
lys
ys
∗
is_list_enc_ref
lzs
zs
∗
⌜
(
x
::
xs
)
=
ys
++
zs
⌝
}}}.
Proof
.
iIntros
(
Φ
)
"[Hhd Heq] HΦ"
.
iDestruct
"Heq"
as
%(
x
&
xs'
&
->).
iDestruct
"Hhd"
as
(
l
hd
->)
"[Hlxs Hhd]"
.
iDestruct
"Hhd"
as
%
Hhd
.
iIntros
(
Φ
)
"Hhd HΦ"
.
iDestruct
"Hhd"
as
(
l
->)
"Hlxs"
.
wp_lam
.
wp_load
.
wp_apply
(
lhead_spec
(
T
:
=
Z
)
hd
)=>
//.
wp_apply
(
lhead_spec
(
T
:
=
Z
))=>
//.
iIntros
"_"
.
wp_pures
.
wp_apply
(
lcons_spec
(
T
:
=
Z
)
_
[])=>
//.
iIntros
(
ys
Hys
).
iIntros
(
_
).
wp_load
.
wp_apply
(
ltail_spec
(
T
:
=
Z
))=>
//.
iIntros
(
zs
Hzs
).
iIntros
(
_
).
wp_alloc
lzs
as
"Hlzs"
.
wp_alloc
lys
as
"Hlys"
.
wp_pures
.
iApply
"HΦ"
.
iFrame
.
rewrite
/
is_list_enc_ref
.
iSplitL
"Hlxs"
.
eauto
10
with
iFrame
.
iSplitL
"Hlys"
.
eauto
10
with
iFrame
.
...
...
@@ -74,8 +71,8 @@ Section ListSortExample.
eauto
.
Qed
.
Definition
compare_vals
:
val
:
=
λ
:
"x"
"y"
,
"x"
≤
"y"
.
Definition
compare_vals
:
val
:
=
λ
:
"x"
"y"
,
"x"
≤
"y"
.
Lemma
compare_vals_spec
(
x
y
:
Z
)
:
{{{
True
%
I
}}}
...
...
@@ -88,6 +85,24 @@ Section ListSortExample.
by
iApply
"HΦ"
.
Qed
.
Definition
my_lemma
(
xs
ys
zs
:
list
Z
)
y
z
:
ordered
Z
.
le
xs
->
ordered
Z
.
le
(
y
::
ys
)
->
ordered
Z
.
le
(
z
::
zs
)
->
permutation
xs
(
ys
++(
z
::
zs
))
->
y
≤
z
->
ordered
Z
.
le
(
y
::
xs
).
Proof
.
Admitted
.
Definition
my_lemma2
(
xs
ys
zs
:
list
Z
)
y
z
:
ordered
Z
.
le
xs
->
ordered
Z
.
le
(
y
::
ys
)
->
ordered
Z
.
le
(
z
::
zs
)
->
permutation
xs
((
y
::
ys
)++
zs
)
->
z
≤
y
->
ordered
Z
.
le
(
z
::
xs
).
Proof
.
Admitted
.
Definition
lmerge
:
val
:
=
rec
:
"go"
"hys"
"hzs"
:
=
if
:
llength
"hys"
=
#
0
...
...
@@ -100,15 +115,16 @@ Section ListSortExample.
then
lcons
"y"
(
"go"
(
ltail
"hys"
)
"hzs"
)
else
lcons
"z"
(
"go"
"hys"
(
ltail
"hzs"
)).
Lemma
lmerge_spec
hdy
hdz
(
ys
zs
:
list
Z
)
:
{{{
⌜
sorted_hd
hdy
ys
⌝
∗
⌜
sorted_hd
hdz
zs
⌝
}}}
lmerge
hdy
hdz
{{{
hxs
xs
,
RET
hxs
;
⌜
sorted_of_hd
hxs
xs
(
ys
++
zs
)
⌝
}}}.
(* Change to just functional version of merge sort *)
Lemma
lmerge_spec
(
ys
zs
:
list
Z
)
:
{{{
⌜
sorted
ys
⌝
∗
⌜
sorted
zs
⌝
}}}
lmerge
(
encode
ys
)
(
encode
zs
)
{{{
xs
,
RET
(
encode
xs
)
;
⌜
sorted_of
xs
(
ys
++
zs
)
⌝
}}}.
Proof
.
revert
hdy
hdz
ys
zs
.
revert
ys
zs
.
iL
ö
b
as
"IH"
.
iIntros
(
hdy
hdz
ys
zs
).
iIntros
(
Φ
[
[
Hhdy
Hhdy_sorted
]
[
Hhdz
Hhdz
_sorted
]
]
)
"HΦ"
.
iIntros
(
ys
zs
).
iIntros
(
Φ
[
ys_sorted
zs
_sorted
])
"HΦ"
.
wp_lam
.
wp_pures
.
wp_apply
(
llength_spec
(
T
:
=
Z
))=>
//.
...
...
@@ -116,74 +132,80 @@ Section ListSortExample.
destruct
ys
as
[|
y
ys
].
{
simpl
.
wp_pures
.
iApply
(
"HΦ"
$!
hdz
zs
)
=>
//.
simpl
.
iApply
(
"HΦ"
$!
zs
)
=>
//.
simpl
.
iSplit
=>
//.
iPureIntro
.
rewrite
/
sorted_of
.
split
=>
//.
apply
permutation_refl
.
iPureIntro
.
apply
permutation_refl
.
}
wp_pures
.
wp_bind
(
llength
_
).
wp_apply
(
llength_spec
(
T
:
=
Z
))=>
//.
iIntros
"_"
.
destruct
zs
as
[|
z
zs
].
{
simpl
.
wp_pures
.
iApply
(
"HΦ"
$!
hdy
(
y
::
ys
))
=>
//.
repeat
rewrite
app_nil_r
.
simpl
.
iApply
(
"HΦ"
$!(
y
::
ys
))
=>
//.
repeat
rewrite
app_nil_r
.
simpl
.
iSplit
=>
//.
iPureIntro
.
rewrite
/
sorted_of
.
split
=>
//.
apply
permutation_refl
.
iPureIntro
.
apply
permutation_refl
.
}
wp_pures
.
wp_apply
(
lhead_spec
(
T
:
=
Z
))=>
//.
iIntros
"_"
.
wp_apply
(
lhead_spec
(
T
:
=
Z
))=>
//.
iIntros
"_"
.
wp_pures
.
wp_apply
(
compare_vals_spec
)=>
//.
iIntros
"_"
.
wp_pures
.
destruct
(
bool_decide
(
y
≤
z
)).
-
wp_pures
.
wp_apply
(
ltail_spec
(
T
:
=
Z
))=>
//.
iIntros
(
hdy'
Hhdy'
).
wp_apply
(
"IH"
)=>
//.
rewrite
/
sorted_hd
.
iSplit
;
iSplit
=>//.
iPureIntro
.
eapply
ordered_cons
=>
//.
iIntros
(
hxs
xs
[
Hhxs
[
Hhxs_sorted
Hhxs_of
]]).
iIntros
(
_
).
wp_apply
(
"IH"
)=>
//.
rewrite
/
sorted
.
iSplit
=>//.
iPureIntro
.
eapply
ordered_cons
=>
//.
iIntros
(
xs
[
Hhxs_sorted
Hhxs_of
]).
wp_apply
(
lcons_spec
(
T
:
=
Z
))=>
//.
iIntros
(
hd
Hhd
).
iApply
(
"HΦ"
).
iIntros
(
_
).
iApply
(
"HΦ"
$!(
y
::
xs
)
).
iPureIntro
.
admit
.
simpl
.
assert
(
y
≤
z
).
admit
.
split
=>
//.
+
eapply
my_lemma
=>
//.
+
by
apply
permutation_cons
.
-
wp_pures
.
wp_apply
(
ltail_spec
(
T
:
=
Z
))=>
//.
iIntros
(
hdz'
Hhdz'
).
wp_apply
(
"IH"
)=>
//.
rewrite
/
sorted
_hd
.
iSplit
;
iSplit
=>//.
iPureIntro
.
eapply
ordered_cons
=>
//.
iIntros
(
hxs
xs
[
Hh
xs
[
Hhxs_sorted
Hhxs_of
]
]
).
iIntros
(
_
).
wp_apply
(
"IH"
)=>
//.
rewrite
/
sorted
.
iSplit
=>//.
iPureIntro
.
eapply
ordered_cons
=>
//.
iIntros
(
xs
[
Hhxs_sorted
Hhxs_of
]).
wp_apply
(
lcons_spec
(
T
:
=
Z
))=>
//.
iIntros
(
hd
Hhd
).
iIntros
(
_
).
iApply
(
"HΦ"
).
iPureIntro
.
admit
.
simpl
.
assert
(
z
≤
y
).
admit
.
split
=>
//.
+
eapply
(
my_lemma2
xs
ys
zs
)=>
//.
+
apply
permutation_sym
.
admit
.
Admitted
.
Definition
lmerge_ref
:
val
:
=
λ
:
"lxs"
"hys"
"hzs"
,
"lxs"
<-
lmerge
"hys"
"hzs"
.
Lemma
lmerge_ref_spec
(
ldx
hdy
hdz
:
val
)
xs
ys
zs
ys'
zs'
:
{{{
⌜
sorted_of
_hd
hdy
ys'
ys
⌝
∗
⌜
sorted_of
_hd
hdz
zs'
zs
⌝
∗
Lemma
lmerge_ref_spec
ldx
xs
ys
zs
ys'
zs'
:
{{{
⌜
sorted_of
ys'
ys
⌝
∗
⌜
sorted_of
zs'
zs
⌝
∗
⌜
xs
=
ys
++
zs
⌝
∗
is_list_enc_ref
ldx
xs
}}}
lmerge_ref
ldx
hdy
hdz
lmerge_ref
ldx
(
encode
ys'
)
(
encode
zs'
)
{{{
xs'
,
RET
#()
;
sorted_of_ref
ldx
xs'
xs
}}}.
Proof
.
Proof
.
iIntros
(
Φ
)
"HPre HΦ"
.
iDestruct
"HPre"
as
(
[
Hhdy
[
Hys_sorted
Hys_perm
]
]
[
Hhdz
[
Hzs_sorted
Hzs_perm
]
]
Heq
)
iDestruct
"HPre"
as
([
Hys_sorted
Hys_perm
]
[
Hzs_sorted
Hzs_perm
]
Heq
)
"Hldx"
.
iDestruct
"Hldx"
as
(
l
hd
->)
"
[Hl
Hldx
]
"
.
iDestruct
"Hldx"
as
(
l
->)
"Hldx"
.
wp_lam
.
wp_apply
(
lmerge_spec
hdy
hdz
ys'
zs'
)=>
//.
iIntros
(
hxs
xs'
)
"Hhxs"
.
iDestruct
"Hhxs"
as
%(
Hhxs
&
Hhys
&
HHzs
).
wp_apply
(
lmerge_spec
ys'
zs'
)=>
//.
iIntros
(
xs'
)
"Hhxs"
.
iDestruct
"Hhxs"
as
%(
Hhxs
_sorted
&
HHxs_perm
).
wp_store
.
iApply
(
"HΦ"
$!
xs'
).
rewrite
/
sorted_of_ref
/
sorted_of
/
is_list_enc_ref
.
...
...
@@ -213,9 +235,9 @@ Section ListSortExample.
send
"c"
#
Right
"xs"
.
Lemma
list_sort_service_spec
γ
c
xs
:
{{{
⟦
c
@
Right
:
(<?>
hd
@
is_list_enc_ref
hd
xs
,
<!>
hd
'
@
⌜
hd
=
hd
'
⌝
∗
(
∃
ys
:
list
Z
,
sorted_of_ref
hd
'
ys
xs
),
{{{
⟦
c
@
Right
:
(<?>
l
@
is_list_enc_ref
l
xs
,
<!>
l
'
@
⌜
l
=
l
'
⌝
∗
(
∃
ys
:
list
Z
,
sorted_of_ref
l
'
ys
xs
),
TEnd
)
⟧
{
N
,
γ
}
}}}
list_sort_service
c
{{{
RET
#()
;
⟦
c
@
Right
:
TEnd
⟧
{
N
,
γ
}
}}}.
...
...
@@ -228,9 +250,8 @@ Section ListSortExample.
iIntros
(
v
)
"Hstr"
.
iDestruct
"Hstr"
as
"[Hstr HP]"
.
wp_pures
.
iDestruct
"HP"
as
(
l
hd'
->)
"
[
Hl
Hhd']
"
.
iDestruct
"HP"
as
(
l
->)
"Hl"
.
wp_load
.
iDestruct
(
"Hhd'"
)
as
%
Hhd'
.
destruct
xs
.
+
rewrite
/
is_list_ref
.
wp_apply
(
llength_spec
(
T
:
=
Z
))=>
//.
...
...
@@ -270,11 +291,6 @@ Section ListSortExample.
rewrite
/(
is_list_enc_ref
(
encode
#
l
)).
eauto
10
with
iFrame
.
iIntros
(
hdy
hdz
ys
zs
)
"(Hhdx & Hhdy & Hhdz & Heq)"
.
wp_pures
.
iDestruct
"Hhdy"
as
(
ly
hdy'
->)
"[Hhdy Hhdys]"
.
iDestruct
"Hhdys"
as
%
Hhdys
.
iDestruct
"Hhdz"
as
(
lz
hdz'
->)
"[Hhdz Hhdzs]"
.
iDestruct
"Hhdzs"
as
%
Hhdzs
.
iDestruct
"Heq"
as
%->.
wp_apply
(
new_chan_st_spec
N
(<!>
hd
@
is_list_enc_ref
hd
ys
,
<?>
hd'
@
(
⌜
hd
=
hd'
⌝
∗
∃
ys'
,
...
...
@@ -290,33 +306,26 @@ Section ListSortExample.
wp_apply
(
wp_fork
with
"[Hstrz]"
).
{
iApply
(
"IH"
with
"Hstrz"
).
iNext
.
iNext
.
iIntros
"Hstrz"
.
done
.
}
wp_apply
(
send_st_spec
N
val
with
"[Hhdy Hstly]"
).
iFrame
.
{
rewrite
/
is_list_enc_ref
.
iExists
_
,
_
.
iSplit
=>
//.
iFrame
.
eauto
10
with
iFrame
.
}
iIntros
"Hstly"
.
wp_apply
(
send_st_spec
N
val
with
"[Hhdz Hstlz]"
).
iFrame
.
{
rewrite
/
is_list_enc_ref
.
iExists
_
,
_
.
iSplit
=>
//.
iFrame
.
eauto
10
with
iFrame
.
}
iIntros
"Hstlz"
.
wp_apply
(
recv_st_spec
N
val
with
"Hstly"
).
iIntros
(
hdy
)
"[Hstly Hys']"
.
iIntros
(
ly'
)
"[Hstly Hys']"
.
iDestruct
"Hys'"
as
(<-)
"Hys'_sorted"
.
iDestruct
"Hys'_sorted"
as
(
ys'
)
"[Hys' Hys'_sorted]"
.
iDestruct
"Hys'_sorted"
as
%
Hys'_sorted
.
iDestruct
"Hys'"
as
(
ly'
Hy'
->)
"[Hlys' Hys']"
.
iDestruct
"Hys'"
as
%
Hys'
.
iDestruct
"Hys'"
as
(
ly'
->)
"Hlys'"
.
wp_apply
(
recv_st_spec
N
val
with
"Hstlz"
).
iIntros
(
hdz
)
"[Hstlz Hzs']"
.
iIntros
(
lz'
)
"[Hstlz Hzs']"
.
iDestruct
"Hzs'"
as
(<-)
"Hzs'_sorted"
.
iDestruct
"Hzs'_sorted"
as
(
zs'
)
"[Hzs' Hzs'_sorted]"
.
iDestruct
"Hzs'_sorted"
as
%
Hzs'_sorted
.
iDestruct
"Hzs'"
as
(
lz'
Hz'
->)
"
[
Hlzs'
Hzs']
"
.
iDestruct
"H
zs'
"
as
%
H
zs'
.
iDestruct
"Hzs'"
as
(
lz'
->)
"Hlzs'"
.
iDestruct
"H
eq
"
as
%
H
eq
.
wp_load
.
wp_load
.
wp_apply
(
lmerge_ref_spec
with
"[Hhdx]"
)=>
//.
iFrame
.
rewrite
/
sorted_hd
.
eauto
10
with
iFrame
.
wp_apply
(
lmerge_ref_spec
(
encode
l
)
_
ys
zs
with
"[Hhdx]"
)=>
//.
eauto
10
with
iFrame
.
iIntros
(
xs'
)
"Hxs'_sorted"
.
wp_apply
(
send_st_spec
N
val
with
"[Hstr Hxs'_sorted]"
).
iFrame
.
eauto
10
with
iFrame
.
...
...
@@ -331,10 +340,10 @@ Section ListSortExample.
send
"c"
#
Left
"xs"
;;
recv
"c"
#
Left
.
Lemma
list_sort_spec
hd
xs
:
{{{
is_list_enc_ref
hd
xs
}}}
list_sort
hd
{{{
ys
,
RET
hd
;
sorted_of_ref
hd
ys
xs
}}}.
Lemma
list_sort_spec
l
xs
:
{{{
is_list_enc_ref
l
xs
}}}