-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy patharray.v
275 lines (257 loc) · 10.5 KB
/
array.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
From iris.proofmode Require Import tactics.
From iris.program_logic Require Export weakestpre.
From iris.heap_lang Require Export derived_laws.
From iris.heap_lang Require Import proofmode notation.
From iris.prelude Require Import options.
(** Provides some array utilities:
* [array_free], to deallocate an entire array in one go.
* [array_copy_to], a function which copies to an array in-place.
* Using [array_copy_to] we also implement [array_clone], which allocates a fresh
array and copies to it.
* [array_init], to create and initialize an array with a given
function. Specifically, [array_init n f] creates a new array of size
[n] in which the [i]th element is initialized with [f #i]
*)
Definition array_free : val :=
rec: "freeN" "ptr" "n" :=
if: "n" ≤ #0 then #()
else Free "ptr";;
"freeN" ("ptr" +ₗ #1) ("n" - #1).
Definition array_copy_to : val :=
rec: "array_copy_to" "dst" "src" "n" :=
if: "n" ≤ #0 then #()
else "dst" <- !"src";;
"array_copy_to" ("dst" +ₗ #1) ("src" +ₗ #1) ("n" - #1).
Definition array_clone : val :=
λ: "src" "n",
let: "dst" := AllocN "n" #() in
array_copy_to "dst" "src" "n";;
"dst".
(* [array_init_loop src i n f] initializes elements
[i], [i+1], ..., [n] of the array [src] to
[f #i], [f #(i+1)], ..., [f #n] *)
Local Definition array_init_loop : val :=
rec: "loop" "src" "i" "n" "f" :=
if: "i" = "n" then #()
else "src" +ₗ "i" <- "f" "i";;
"loop" "src" ("i" + #1) "n" "f".
Definition array_init : val :=
λ: "n" "f",
let: "src" := AllocN "n" #() in
array_init_loop "src" #0 "n" "f";;
"src".
Section proof.
Context `{!heapG Σ}.
Lemma twp_array_free s E l vs (n : Z) :
n = length vs →
[[{ l ↦∗ vs }]] array_free #l #n @ s; E [[{ RET #(); True }]].
Proof.
iIntros (Hlen Φ) "Hl HΦ".
iInduction vs as [|v vs] "IH" forall (l n Hlen); subst n; wp_rec; wp_pures.
{ iApply "HΦ". done. }
iDestruct (array_cons with "Hl") as "[Hv Hl]".
wp_free. wp_pures. iApply ("IH" with "[] Hl"); eauto with lia.
Qed.
Lemma wp_array_free s E l vs (n : Z) :
n = length vs →
{{{ l ↦∗ vs }}} array_free #l #n @ s; E {{{ RET #(); True }}}.
Proof.
iIntros (? Φ) "H HΦ". iApply (twp_wp_step with "HΦ").
iApply (twp_array_free with "H"); [auto..|]; iIntros "H HΦ". by iApply "HΦ".
Qed.
Lemma twp_array_copy_to stk E (dst src : loc) vdst vsrc dq (n : Z) :
Z.of_nat (length vdst) = n → Z.of_nat (length vsrc) = n →
[[{ dst ↦∗ vdst ∗ src ↦∗{dq} vsrc }]]
array_copy_to #dst #src #n @ stk; E
[[{ RET #(); dst ↦∗ vsrc ∗ src ↦∗{dq} vsrc }]].
Proof.
iIntros (Hvdst Hvsrc Φ) "[Hdst Hsrc] HΦ".
iInduction vdst as [|v1 vdst] "IH" forall (n dst src vsrc Hvdst Hvsrc);
destruct vsrc as [|v2 vsrc]; simplify_eq/=; try lia; wp_rec; wp_pures.
{ iApply "HΦ". auto with iFrame. }
iDestruct (array_cons with "Hdst") as "[Hv1 Hdst]".
iDestruct (array_cons with "Hsrc") as "[Hv2 Hsrc]".
wp_load; wp_store.
wp_smart_apply ("IH" with "[%] [%] Hdst Hsrc"); [ lia .. | ].
iIntros "[Hvdst Hvsrc]".
iApply "HΦ"; by iFrame.
Qed.
Lemma wp_array_copy_to stk E (dst src : loc) vdst vsrc dq (n : Z) :
Z.of_nat (length vdst) = n → Z.of_nat (length vsrc) = n →
{{{ dst ↦∗ vdst ∗ src ↦∗{dq} vsrc }}}
array_copy_to #dst #src #n @ stk; E
{{{ RET #(); dst ↦∗ vsrc ∗ src ↦∗{dq} vsrc }}}.
Proof.
iIntros (? ? Φ) "H HΦ". iApply (twp_wp_step with "HΦ").
iApply (twp_array_copy_to with "H"); [auto..|]; iIntros "H HΦ". by iApply "HΦ".
Qed.
Lemma twp_array_clone stk E l dq vl n :
Z.of_nat (length vl) = n →
(0 < n)%Z →
[[{ l ↦∗{dq} vl }]]
array_clone #l #n @ stk; E
[[{ l', RET #l'; l' ↦∗ vl ∗ l ↦∗{dq} vl }]].
Proof.
iIntros (Hvl Hn Φ) "Hvl HΦ".
wp_lam.
wp_alloc dst as "Hdst"; first by auto.
wp_smart_apply (twp_array_copy_to with "[$Hdst $Hvl]").
- rewrite replicate_length Z2Nat.id; lia.
- auto.
- iIntros "[Hdst Hl]".
wp_pures.
iApply "HΦ"; by iFrame.
Qed.
Lemma wp_array_clone stk E l dq vl n :
Z.of_nat (length vl) = n →
(0 < n)%Z →
{{{ l ↦∗{dq} vl }}}
array_clone #l #n @ stk; E
{{{ l', RET #l'; l' ↦∗ vl ∗ l ↦∗{dq} vl }}}.
Proof.
iIntros (? ? Φ) "H HΦ". iApply (twp_wp_step with "HΦ").
iApply (twp_array_clone with "H"); [auto..|]; iIntros (l') "H HΦ". by iApply "HΦ".
Qed.
Section array_init.
Context (Q : nat → val → iProp Σ).
Implicit Types (f v : val) (i j : nat).
Local Lemma wp_array_init_loop stk E l i n k f :
n = Z.of_nat (i + k) →
{{{
(l +ₗ i) ↦∗ replicate k #() ∗
[∗ list] j ∈ seq i k, WP f #(j : nat) @ stk; E {{ Q j }}
}}}
array_init_loop #l #i #n f @ stk; E
{{{ vs, RET #();
⌜ length vs = k ⌝ ∗ (l +ₗ i) ↦∗ vs ∗ [∗ list] j↦v ∈ vs, Q (i + j) v
}}}.
Proof.
iIntros (Hn Φ) "[Hl Hf] HΦ".
iInduction k as [|k] "IH" forall (i Hn); simplify_eq/=; wp_rec; wp_pures.
{ rewrite bool_decide_eq_true_2; last (repeat f_equal; lia).
wp_pures. iApply ("HΦ" $! []). auto. }
rewrite bool_decide_eq_false_2; last naive_solver lia.
iDestruct (array_cons with "Hl") as "[Hl HSl]".
iDestruct "Hf" as "[Hf HSf]".
wp_smart_apply (wp_wand with "Hf"); iIntros (v) "Hv". wp_store. wp_pures.
rewrite Z.add_1_r -Nat2Z.inj_succ.
iApply ("IH" with "[%] [HSl] HSf"); first lia.
{ by rewrite loc_add_assoc Z.add_1_r -Nat2Z.inj_succ. }
iIntros "!>" (vs). iDestruct 1 as (<-) "[HSl Hvs]".
iApply ("HΦ" $! (v :: vs)). iSplit; [naive_solver|]. iSplitL "Hl HSl".
- iFrame "Hl". by rewrite loc_add_assoc Z.add_1_r -Nat2Z.inj_succ.
- iEval (rewrite /= Nat.add_0_r; setoid_rewrite Nat.add_succ_r). iFrame.
Qed.
Local Lemma twp_array_init_loop stk E l i n k f :
n = Z.of_nat (i + k) →
[[{
(l +ₗ i) ↦∗ replicate k #() ∗
[∗ list] j ∈ seq i k, WP f #(j : nat) @ stk; E [{ Q j }]
}]]
array_init_loop #l #i #n f @ stk; E
[[{ vs, RET #();
⌜ length vs = k ⌝ ∗ (l +ₗ i) ↦∗ vs ∗ [∗ list] j↦v ∈ vs, Q (i + j) v
}]].
Proof.
iIntros (Hn Φ) "[Hl Hf] HΦ".
iInduction k as [|k] "IH" forall (i Hn); simplify_eq/=; wp_rec; wp_pures.
{ rewrite bool_decide_eq_true_2; last (repeat f_equal; lia).
wp_pures. iApply ("HΦ" $! []). auto. }
rewrite bool_decide_eq_false_2; last naive_solver lia.
iDestruct (array_cons with "Hl") as "[Hl HSl]".
iDestruct "Hf" as "[Hf HSf]".
wp_smart_apply (twp_wand with "Hf"); iIntros (v) "Hv". wp_store. wp_pures.
rewrite Z.add_1_r -Nat2Z.inj_succ.
iApply ("IH" with "[%] [HSl] HSf"); first lia.
{ by rewrite loc_add_assoc Z.add_1_r -Nat2Z.inj_succ. }
iIntros (vs). iDestruct 1 as (<-) "[HSl Hvs]".
iApply ("HΦ" $! (v :: vs)). iSplit; [naive_solver|]. iSplitL "Hl HSl".
- iFrame "Hl". by rewrite loc_add_assoc Z.add_1_r -Nat2Z.inj_succ.
- iEval (rewrite /= Nat.add_0_r; setoid_rewrite Nat.add_succ_r). iFrame.
Qed.
Lemma wp_array_init stk E n f :
(0 < n)%Z →
{{{
[∗ list] i ∈ seq 0 (Z.to_nat n), WP f #(i : nat) @ stk; E {{ Q i }}
}}}
array_init #n f @ stk; E
{{{ l vs, RET #l;
⌜Z.of_nat (length vs) = n⌝ ∗ l ↦∗ vs ∗ [∗ list] k↦v ∈ vs, Q k v
}}}.
Proof.
iIntros (Hn Φ) "Hf HΦ". wp_lam. wp_alloc l as "Hl"; first done.
wp_smart_apply (wp_array_init_loop _ _ _ 0 n (Z.to_nat n) with "[Hl $Hf] [HΦ]").
{ by rewrite /= Z2Nat.id; last lia. }
{ by rewrite loc_add_0. }
iIntros "!>" (vs). iDestruct 1 as (Hlen) "[Hl Hvs]". wp_pures.
iApply ("HΦ" $! _ vs). iModIntro. iSplit.
{ iPureIntro. by rewrite Hlen Z2Nat.id; last lia. }
rewrite loc_add_0. iFrame.
Qed.
Lemma twp_array_init stk E n f :
(0 < n)%Z →
[[{
[∗ list] i ∈ seq 0 (Z.to_nat n), WP f #(i : nat) @ stk; E [{ Q i }]
}]]
array_init #n f @ stk; E
[[{ l vs, RET #l;
⌜Z.of_nat (length vs) = n⌝ ∗ l ↦∗ vs ∗ [∗ list] k↦v ∈ vs, Q k v
}]].
Proof.
iIntros (Hn Φ) "Hf HΦ". wp_lam. wp_alloc l as "Hl"; first done.
wp_smart_apply (twp_array_init_loop _ _ _ 0 n (Z.to_nat n) with "[Hl $Hf] [HΦ]").
{ by rewrite /= Z2Nat.id; last lia. }
{ by rewrite loc_add_0. }
iIntros (vs). iDestruct 1 as (Hlen) "[Hl Hvs]". wp_pures.
iApply ("HΦ" $! _ vs). iModIntro. iSplit.
{ iPureIntro. by rewrite Hlen Z2Nat.id; last lia. }
rewrite loc_add_0. iFrame.
Qed.
End array_init.
Section array_init_fmap.
Context {A} (g : A → val) (Q : nat → A → iProp Σ).
Implicit Types (xs : list A) (f : val).
Local Lemma big_sepL_exists_eq vs :
([∗ list] k↦v ∈ vs, ∃ x, ⌜v = g x⌝ ∗ Q k x) -∗
∃ xs, ⌜ vs = g <$> xs ⌝ ∗ [∗ list] k↦x ∈ xs, Q k x.
Proof.
iIntros "Hvs". iInduction vs as [|v vs] "IH" forall (Q); simpl.
{ iExists []. by auto. }
iDestruct "Hvs" as "[Hv Hvs]"; iDestruct "Hv" as (x ->) "Hv".
iDestruct ("IH" with "Hvs") as (xs ->) "Hxs".
iExists (x :: xs). by iFrame.
Qed.
Lemma wp_array_init_fmap stk E n f :
(0 < n)%Z →
{{{
[∗ list] i ∈ seq 0 (Z.to_nat n),
WP f #(i : nat) @ stk; E {{ v, ∃ x, ⌜v = g x⌝ ∗ Q i x }}
}}}
array_init #n f @ stk; E
{{{ l xs, RET #l;
⌜Z.of_nat (length xs) = n⌝ ∗ l ↦∗ (g <$> xs) ∗ [∗ list] k↦x ∈ xs, Q k x
}}}.
Proof.
iIntros (Hn Φ) "Hf HΦ". iApply (wp_array_init with "Hf"); first done.
iIntros "!>" (l vs). iDestruct 1 as (<-) "[Hl Hvs]".
iDestruct (big_sepL_exists_eq with "Hvs") as (xs ->) "Hxs".
iApply "HΦ". iFrame "Hl Hxs". by rewrite fmap_length.
Qed.
Lemma twp_array_init_fmap stk E n f :
(0 < n)%Z →
[[{
[∗ list] i ∈ seq 0 (Z.to_nat n),
WP f #(i : nat) @ stk; E [{ v, ∃ x, ⌜v = g x⌝ ∗ Q i x }]
}]]
array_init #n f @ stk; E
[[{ l xs, RET #l;
⌜Z.of_nat (length xs) = n⌝ ∗ l ↦∗ (g <$> xs) ∗ [∗ list] k↦x ∈ xs, Q k x
}]].
Proof.
iIntros (Hn Φ) "Hf HΦ". iApply (twp_array_init with "Hf"); first done.
iIntros (l vs). iDestruct 1 as (<-) "[Hl Hvs]".
iDestruct (big_sepL_exists_eq with "Hvs") as (xs ->) "Hxs".
iApply "HΦ". iFrame "Hl Hxs". by rewrite fmap_length.
Qed.
End array_init_fmap.
End proof.