-
Notifications
You must be signed in to change notification settings - Fork 108
/
Copy pathMonad_Lists.thy
707 lines (616 loc) · 23.9 KB
/
Monad_Lists.thy
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
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
698
699
700
701
702
703
704
705
706
707
(*
* Copyright 2023, Proofcraft Pty Ltd
* Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
*
* SPDX-License-Identifier: BSD-2-Clause
*)
(* Monadic functions over lists: sequence, mapM, filter, etc
Definitions, equations, Hoare logic and no_fail/empty_fail setup. *)
theory Monad_Lists
imports
Monads.Nondet_In_Monad
Monads.Nondet_Det
Monads.Nondet_Empty_Fail
Monads.Nondet_No_Fail
begin
lemma mapME_Cons:
"mapME m (x # xs) = (doE y \<leftarrow> m x; ys \<leftarrow> (mapME m xs); returnOk (y # ys) odE)"
by (simp add: mapME_def sequenceE_def Let_def)
lemma mapME_Nil : "mapME f [] = returnOk []"
unfolding mapME_def by (simp add: sequenceE_def)
lemmas mapME_simps = mapME_Nil mapME_Cons
lemma zipWithM_x_inv':
assumes x: "\<And>x y. m x y \<lbrace>P\<rbrace>"
shows "zipWithM_x m xs ys \<lbrace>P\<rbrace>"
proof (induct xs arbitrary: ys)
case Nil
show ?case
by (simp add: zipWithM_x_def sequence_x_def zipWith_def)
next
case (Cons x xs)
have zipWithM_x_Cons:
"\<And>m x xs y ys. zipWithM_x m (x # xs) (y # ys) = do m x y; zipWithM_x m xs ys od"
by (simp add: zipWithM_x_def sequence_x_def zipWith_def)
have zipWithM_x_Nil:
"\<And>m xs. zipWithM_x m xs [] = return ()"
by (simp add: zipWithM_x_def sequence_x_def zipWith_def)
show ?case
by (cases ys; wpsimp simp: zipWithM_x_Nil zipWithM_x_Cons wp: Cons x)
qed
(* For compatibility with existing proofs. *)
lemma zipWithM_x_inv:
assumes x: "\<And>x y. m x y \<lbrace>P\<rbrace>"
shows "length xs = length ys \<Longrightarrow> zipWithM_x m xs ys \<lbrace>P\<rbrace>"
by (rule zipWithM_x_inv', rule x)
lemma sequence_x_Cons: "\<And>x xs. sequence_x (x # xs) = (x >>= (\<lambda>_. sequence_x xs))"
by (simp add: sequence_x_def)
lemma mapM_Cons: "mapM m (x # xs) = (do y \<leftarrow> m x; ys \<leftarrow> (mapM m xs); return (y # ys) od)"
by (simp add: mapM_def sequence_def Let_def)
lemma mapM_Nil:
"mapM m [] = return []"
by (simp add: mapM_def sequence_def)
lemmas mapM_simps = mapM_Nil mapM_Cons
lemma zipWithM_x_mapM:
"zipWithM_x f as bs = (mapM (case_prod f) (zip as bs) >>= (\<lambda>_. return ()))"
apply (simp add: zipWithM_x_def zipWith_def)
apply (induct ("zip as bs"))
apply (simp add: sequence_x_def mapM_def sequence_def)
apply (simp add: sequence_x_Cons mapM_Cons bind_assoc)
done
lemma mapM_x_mapM:
"mapM_x m l = (mapM m l >>= (\<lambda>x. return ()))"
apply (simp add: mapM_x_def sequence_x_def mapM_def sequence_def)
apply (induct l, simp_all add: Let_def bind_assoc)
done
lemma mapM_x_Nil:
"mapM_x f [] = return ()"
unfolding mapM_x_def sequence_x_def
by simp
lemma sequence_xappend1:
"sequence_x (xs @ [x]) = (sequence_x xs >>= (\<lambda>_. x))"
by (induct xs) (simp add: sequence_x_def, simp add: sequence_x_Cons bind_assoc)
lemma mapM_append_single:
"mapM_x f (xs @ [y]) = (mapM_x f xs >>= (\<lambda>_. f y))"
unfolding mapM_x_def
by (simp add: sequence_xappend1)
lemma mapM_x_Cons:
"mapM_x m (x # xs) = (do m x; mapM_x m xs od)"
by (simp add: mapM_x_def sequence_x_def)
lemma zipWithM_x_mapM_x:
"zipWithM_x f as bs = mapM_x (\<lambda>(x, y). f x y) (zip as bs)"
apply (subst zipWithM_x_mapM)
apply (subst mapM_x_mapM)
apply (rule refl)
done
lemma zipWithM_x_append1:
fixes f :: "'b \<Rightarrow> 'c \<Rightarrow> ('a, unit) nondet_monad"
assumes ls: "length xs = length ys"
shows "(zipWithM_x f (xs @ [x]) (ys @ [y])) = (zipWithM_x f xs ys >>= (\<lambda>_. f x y))"
unfolding zipWithM_x_def zipWith_def
by (subst zip_append [OF ls], simp, rule sequence_xappend1)
lemma zipWithM_x_Cons:
assumes ls: "length xs = length ys"
shows "(zipWithM_x f (x # xs) (y # ys)) = (f x y >>= (\<lambda>_. zipWithM_x f xs ys))"
unfolding zipWithM_x_def zipWith_def
by (simp, rule sequence_x_Cons)
lemma mapME_x_map_simp:
"mapME_x m (map f xs) = mapME_x (m o f) xs"
by (simp add: mapME_x_def sequenceE_x_def)
lemma mapM_return:
"mapM (\<lambda>x. return (f x)) xs = return (map f xs)"
apply (induct xs)
apply (simp add: mapM_def sequence_def)
apply (simp add: mapM_Cons)
done
lemma liftM_return [simp]:
"liftM f (return x) = return (f x)"
by (simp add: liftM_def)
lemma mapM_x_return :
"mapM_x (\<lambda>_. return v) xs = return v"
by (induct xs) (auto simp: mapM_x_Nil mapM_x_Cons)
lemma bind_comm_mapM_comm:
assumes bind_comm:
"\<And>n z. do x \<leftarrow> a; y \<leftarrow> b z; (n x y :: ('a, 's) nondet_monad) od =
do y \<leftarrow> b z; x \<leftarrow> a; n x y od"
shows "\<And>n'. do x \<leftarrow> a; ys \<leftarrow> mapM b zs; (n' x ys :: ('a, 's) nondet_monad) od =
do ys \<leftarrow> mapM b zs; x \<leftarrow> a; n' x ys od"
proof (induct zs)
case Nil
thus ?case
by (simp add: mapM_def sequence_def)
next
case (Cons z zs')
thus ?case
by (clarsimp simp: mapM_Cons bind_assoc bind_comm intro!: bind_cong [OF refl])
qed
lemma liftE_handle :
"(liftE f <handle> g) = liftE f"
by (simp add: handleE_def handleE'_def liftE_def)
lemma mapM_empty:
"mapM f [] = return []"
unfolding mapM_def
by (simp add: sequence_def)
lemma mapM_append:
"mapM f (xs @ ys) =
(do x \<leftarrow> mapM f xs;
y \<leftarrow> mapM f ys;
return (x @ y)
od)"
proof (induct xs)
case Nil
thus ?case by (simp add: mapM_empty)
next
case (Cons x xs)
show ?case
by (simp add: mapM_Cons bind_assoc Cons.hyps)
qed
lemma mapM_x_append: (* FIXME: remove extra return, fix proofs *)
"mapM_x f (xs @ ys) =
(do x \<leftarrow> mapM_x f xs;
y \<leftarrow> mapM_x f ys;
return ()
od)"
by (simp add: mapM_x_mapM mapM_append bind_assoc)
(* FIXME: duplicate, but mapM_x_append has an extra useless return *)
lemma mapM_x_append2:
"mapM_x f (xs @ ys) = do mapM_x f xs; mapM_x f ys od"
apply (simp add: mapM_x_def sequence_x_def)
apply (induct xs)
apply simp
apply (simp add: bind_assoc)
done
lemma mapM_singleton:
"mapM f [x] = do r \<leftarrow> f x; return [r] od"
by (simp add: mapM_def sequence_def)
lemma mapM_x_singleton:
"mapM_x f [x] = f x"
by (simp add: mapM_x_mapM mapM_singleton)
lemma mapME_x_sequenceE:
"mapME_x f xs \<equiv> doE _ \<leftarrow> sequenceE (map f xs); returnOk () odE"
apply (induct xs, simp_all add: mapME_x_def sequenceE_def sequenceE_x_def)
apply (simp add: Let_def bindE_assoc)
done
lemma sequenceE_Cons:
"sequenceE (x # xs) = (doE v \<leftarrow> x; vs \<leftarrow> sequenceE xs; returnOk (v # vs) odE)"
by (simp add: sequenceE_def Let_def)
lemma zipWithM_Nil [simp]:
"zipWithM f xs [] = return []"
by (simp add: zipWithM_def zipWith_def sequence_def)
lemma zipWithM_One:
"zipWithM f (x#xs) [a] = (do z \<leftarrow> f x a; return [z] od)"
by (simp add: zipWithM_def zipWith_def sequence_def)
lemma zipWithM_x_Nil[simp]:
"zipWithM_x f xs [] = return ()"
by (simp add: zipWithM_x_def zipWith_def sequence_x_def)
lemma zipWithM_x_One:
"zipWithM_x f (x#xs) [a] = f x a"
by (simp add: zipWithM_x_def zipWith_def sequence_x_def)
lemma mapM_last_Cons:
"\<lbrakk> xs = [] \<Longrightarrow> g v = y;
xs \<noteq> [] \<Longrightarrow> do x \<leftarrow> f (last xs); return (g x) od
= do x \<leftarrow> f (last xs); return y od \<rbrakk> \<Longrightarrow>
do ys \<leftarrow> mapM f xs; return (g (last (v # ys))) od = do mapM_x f xs; return y od"
apply (cases "xs = []")
apply (simp add: mapM_x_Nil mapM_Nil)
apply (simp add: mapM_x_mapM)
apply (subst append_butlast_last_id[symmetric], assumption,
subst mapM_append)+
apply (simp add: bind_assoc mapM_Cons mapM_Nil)
done
lemma map_length_cong:
"\<lbrakk> length xs = length ys; \<And>x y. (x, y) \<in> set (zip xs ys) \<Longrightarrow> f x = g y \<rbrakk>
\<Longrightarrow> map f xs = map g ys"
apply atomize
apply (erule rev_mp, erule list_induct2)
apply auto
done
lemma mapM_length_cong:
"\<lbrakk> length xs = length ys; \<And>x y. (x, y) \<in> set (zip xs ys) \<Longrightarrow> f x = g y \<rbrakk>
\<Longrightarrow> mapM f xs = mapM g ys"
by (simp add: mapM_def map_length_cong)
(* FIXME: duplicate *)
lemma zipWithM_mapM:
"zipWithM f xs ys = mapM (case_prod f) (zip xs ys)"
by (simp add: zipWithM_def zipWith_def mapM_def)
lemma zip_take_triv2:
"length as \<le> n \<Longrightarrow> zip as (take n bs) = zip as bs"
apply (induct as arbitrary: n bs; simp)
apply (case_tac n; simp)
apply (case_tac bs; simp)
done
lemma zipWithM_If_cut:
"zipWithM (\<lambda>a b. if a < n then f a b else g a b) [0 ..< m] xs
= do ys \<leftarrow> zipWithM f [0 ..< min n m] xs;
zs \<leftarrow> zipWithM g [n ..< m] (drop n xs);
return (ys @ zs)
od"
apply (cases "n < m")
apply (cut_tac i=0 and j=n and k="m - n" in upt_add_eq_append)
apply simp
apply (simp add: zipWithM_mapM)
apply (simp add: zip_append1 mapM_append zip_take_triv2 split_def)
apply (intro bind_cong bind_apply_cong refl mapM_length_cong
fun_cong[OF mapM_length_cong])
apply (clarsimp simp: set_zip)
apply (clarsimp simp: set_zip)
apply (simp add: zipWithM_mapM mapM_Nil)
apply (intro mapM_length_cong refl)
apply (clarsimp simp: set_zip)
done
lemma mapM_liftM_const:
"mapM (\<lambda>x. liftM (\<lambda>y. f x) (g x)) xs = liftM (\<lambda>ys. map f xs) (mapM g xs)"
apply (induct xs)
apply (simp add: mapM_Nil)
apply (simp add: mapM_Cons)
apply (simp add: liftM_def bind_assoc)
done
lemma mapM_discarded:
"mapM f xs >>= (\<lambda>ys. g) = mapM_x f xs >>= (\<lambda>_. g)"
by (simp add: mapM_x_mapM)
lemma mapM_x_map:
"mapM_x f (map g xs) = mapM_x (\<lambda>x. f (g x)) xs"
by (simp add: mapM_x_def o_def)
lemma filterM_append:
"filterM f (xs @ ys) = do
xs' \<leftarrow> filterM f xs;
ys' \<leftarrow> filterM f ys;
return (xs' @ ys')
od"
apply (induct xs)
apply simp
apply (simp add: bind_assoc)
apply (rule ext bind_apply_cong [OF refl])+
apply simp
done
lemma filterM_mapM:
"filterM f xs = do
ys \<leftarrow> mapM (\<lambda>x. do v \<leftarrow> f x; return (x, v) od) xs;
return (map fst (filter snd ys))
od"
apply (induct xs)
apply (simp add: mapM_def sequence_def)
apply (simp add: mapM_Cons bind_assoc)
apply (rule bind_cong [OF refl] bind_apply_cong[OF refl])+
apply simp
done
lemma mapM_gets:
assumes P: "\<And>x. m x = gets (f x)"
shows "mapM m xs = gets (\<lambda>s. map (\<lambda>x. f x s) xs)"
proof (induct xs)
case Nil show ?case
by (simp add: mapM_def sequence_def gets_def get_def bind_def)
next
case (Cons y ys) thus ?case
by (simp add: mapM_Cons P simpler_gets_def return_def bind_def)
qed
lemma mapM_map_simp:
"mapM m (map f xs) = mapM (m \<circ> f) xs"
apply (induct xs)
apply (simp add: mapM_def sequence_def)
apply (simp add: mapM_Cons)
done
lemma filterM_voodoo:
"\<forall>ys. P ys (do zs \<leftarrow> filterM m xs; return (ys @ zs) od)
\<Longrightarrow> P [] (filterM m xs)"
by (drule spec[where x=Nil], simp)
lemma mapME_x_Cons:
"mapME_x f (x # xs) = (doE f x; mapME_x f xs odE)"
by (simp add: mapME_x_def sequenceE_x_def)
lemma liftME_map_mapME:
"liftME (map f) (mapME m xs)
= mapME (liftME f o m) xs"
apply (rule sym)
apply (induct xs)
apply (simp add: liftME_def mapME_Nil)
apply (simp add: mapME_Cons liftME_def bindE_assoc)
done
lemma mapM_x_split_append:
"mapM_x f xs = do _ \<leftarrow> mapM_x f (take n xs); mapM_x f (drop n xs) od"
using mapM_x_append[where f=f and xs="take n xs" and ys="drop n xs"]
by simp
lemma mapME_wp:
assumes x: "\<And>x. x \<in> S \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>_. P\<rbrace>, \<lbrace>\<lambda>_. E\<rbrace>"
shows "set xs \<subseteq> S \<Longrightarrow> \<lbrace>P\<rbrace> mapME f xs \<lbrace>\<lambda>_. P\<rbrace>, \<lbrace>\<lambda>_. E\<rbrace>"
apply (induct xs)
apply (simp add: mapME_def sequenceE_def)
apply wp
apply simp
apply (simp add: mapME_Cons)
apply (wp x|simp)+
done
lemmas mapME_wp' = mapME_wp [OF _ subset_refl]
lemma mapM_x_inv_wp3:
fixes m :: "'b \<Rightarrow> ('a, unit) nondet_monad"
assumes hr: "\<And>a as bs. xs = as @ [a] @ bs \<Longrightarrow>
\<lbrace>\<lambda>s. I s \<and> V as s\<rbrace> m a \<lbrace>\<lambda>r s. I s \<and> V (as @ [a]) s\<rbrace>"
shows "\<lbrace>\<lambda>s. I s \<and> V [] s\<rbrace> mapM_x m xs \<lbrace>\<lambda>rv s. I s \<and> V xs s\<rbrace>"
using hr
proof (induct xs rule: rev_induct)
case Nil thus ?case
by (simp add: mapM_x_Nil)
next
case (snoc x xs)
show ?case
apply (simp add: mapM_append_single)
apply (wp snoc.prems)
apply simp
apply (rule snoc.hyps [OF snoc.prems])
apply simp
apply assumption
done
qed
lemma mapME_x_inv_wp:
assumes x: "\<And>x. \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>rv. P\<rbrace>,\<lbrace>E\<rbrace>"
shows "\<lbrace>P\<rbrace> mapME_x f xs \<lbrace>\<lambda>rv. P\<rbrace>,\<lbrace>E\<rbrace>"
apply (induct xs)
apply (simp add: mapME_x_def sequenceE_x_def)
apply wp
apply (simp add: mapME_x_def sequenceE_x_def)
apply (fold mapME_x_def sequenceE_x_def)
apply wp
apply (rule x)
apply assumption
done
lemma mapM_upd:
assumes "\<And>x rv s s'. (rv,s') \<in> fst (f x s) \<Longrightarrow> x \<in> set xs \<Longrightarrow> (rv, g s') \<in> fst (f x (g s))"
shows "(rv,s') \<in> fst (mapM f xs s) \<Longrightarrow> (rv, g s') \<in> fst (mapM f xs (g s))"
using assms
proof (induct xs arbitrary: rv s s')
case Nil
thus ?case by (simp add: mapM_Nil return_def)
next
case (Cons z zs)
from Cons.prems
show ?case
apply (clarsimp simp: mapM_Cons in_monad)
apply (drule Cons.prems, simp)
apply (rule exI, erule conjI)
apply (erule Cons.hyps)
apply (erule Cons.prems)
apply simp
done
qed
lemma no_fail_mapM_wp:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> no_fail (P x) (f x)"
assumes "\<And>x y. \<lbrakk> x \<in> set xs; y \<in> set xs \<rbrakk> \<Longrightarrow> \<lbrace>P x\<rbrace> f y \<lbrace>\<lambda>_. P x\<rbrace>"
shows "no_fail (\<lambda>s. \<forall>x \<in> set xs. P x s) (mapM f xs)"
using assms
proof (induct xs)
case Nil
thus ?case by (simp add: mapM_empty)
next
case (Cons z zs)
show ?case
apply (clarsimp simp: mapM_Cons)
apply (wp Cons.prems Cons.hyps hoare_vcg_const_Ball_lift|simp)+
done
qed
lemma no_fail_mapM:
"\<forall>x. no_fail \<top> (f x) \<Longrightarrow> no_fail \<top> (mapM f xs)"
apply (induct xs)
apply (simp add: mapM_def sequence_def)
apply (simp add: mapM_Cons)
apply (wp|fastforce)+
done
lemma filterM_preserved:
"\<lbrakk> \<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> m x \<lbrace>\<lambda>rv. P\<rbrace> \<rbrakk>
\<Longrightarrow> \<lbrace>P\<rbrace> filterM m xs \<lbrace>\<lambda>rv. P\<rbrace>"
apply (induct xs)
apply (wp | simp | erule meta_mp | drule meta_spec)+
done
lemma filterM_distinct1:
"\<lbrace>\<top> and K (P \<longrightarrow> distinct xs)\<rbrace> filterM m xs \<lbrace>\<lambda>rv s. (P \<longrightarrow> distinct rv) \<and> set rv \<subseteq> set xs\<rbrace>"
apply (rule hoare_gen_asm, erule rev_mp)
apply (rule rev_induct [where xs=xs])
apply (clarsimp | wp)+
apply (simp add: filterM_append)
apply (erule bind_wp_fwd)
apply (rule bind_wp_fwd, rule hoare_vcg_prop)
apply (wp, clarsimp)
apply blast
done
lemma filterM_subset:
"\<lbrace>\<top>\<rbrace> filterM m xs \<lbrace>\<lambda>rv s. set rv \<subseteq> set xs\<rbrace>"
by (rule hoare_chain, rule filterM_distinct1[where P=False], simp_all)
lemma filterM_all:
"\<lbrakk> \<And>x y. \<lbrakk> x \<in> set xs; y \<in> set xs \<rbrakk> \<Longrightarrow> \<lbrace>P y\<rbrace> m x \<lbrace>\<lambda>rv. P y\<rbrace> \<rbrakk> \<Longrightarrow>
\<lbrace>\<lambda>s. \<forall>x \<in> set xs. P x s\<rbrace> filterM m xs \<lbrace>\<lambda>rv s. \<forall>x \<in> set rv. P x s\<rbrace>"
apply (rule_tac Q'="\<lambda>rv s. set rv \<subseteq> set xs \<and> (\<forall>x \<in> set xs. P x s)"
in hoare_strengthen_post)
apply (wp filterM_subset hoare_vcg_const_Ball_lift filterM_preserved)
apply simp+
apply blast
done
lemma filterM_distinct:
"\<lbrace>K (distinct xs)\<rbrace> filterM m xs \<lbrace>\<lambda>rv s. distinct rv\<rbrace>"
by (rule hoare_chain, rule filterM_distinct1[where P=True], simp_all)
lemma mapM_wp:
assumes x: "\<And>x. x \<in> S \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>rv. P\<rbrace>"
shows "set xs \<subseteq> S \<Longrightarrow> \<lbrace>P\<rbrace> mapM f xs \<lbrace>\<lambda>rv. P\<rbrace>"
apply (induct xs)
apply (simp add: mapM_def sequence_def)
apply (simp add: mapM_Cons)
apply wp
apply (rule x, clarsimp)
apply simp
done
lemma mapM_wp':
assumes x: "\<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>rv. P\<rbrace>"
shows "\<lbrace>P\<rbrace> mapM f xs \<lbrace>\<lambda>rv. P\<rbrace>"
apply (rule mapM_wp)
apply (erule x)
apply simp
done
lemma mapM_gets_the_wp:
"\<lbrace>\<lambda>s. \<forall>vs. map (\<lambda>t. f t s) ts = map Some vs \<longrightarrow> P vs s\<rbrace>
mapM (gets_the \<circ> f) ts
\<lbrace>P\<rbrace>"
unfolding comp_def
proof (induct ts arbitrary: P)
case Nil thus ?case by (wpsimp simp: mapM_Nil)
next
case (Cons x xs) show ?case by (wpsimp simp: mapM_Cons wp: Cons)
qed
lemma mapM_gets_the_wp_nth:
"\<lbrace>\<lambda>s. \<forall>vs. (\<forall>i < length vs. f (ts ! i) s = Some (vs ! i)) \<longrightarrow> P vs s\<rbrace>
mapM (gets_the \<circ> f) ts
\<lbrace>P\<rbrace>"
apply (wpsimp wp: mapM_gets_the_wp)
by (simp add: map_equality_iff)
lemma mapM_gets_the_sp:
"\<lbrace>P\<rbrace> mapM (gets_the \<circ> f) ts \<lbrace>\<lambda>rv s. P s \<and> map (\<lambda>t. f t s) ts = map Some rv\<rbrace>"
by (wpsimp wp: mapM_gets_the_wp simp: comp_def)
lemma mapM_gets_the_sp_nth:
"\<lbrace>P\<rbrace> mapM (gets_the \<circ> f) ts \<lbrace>\<lambda>rv s. P s \<and> (\<forall>i < length rv. f (ts ! i) s = Some (rv ! i))\<rbrace>"
apply (wpsimp wp: mapM_gets_the_wp simp: comp_def)
by (simp add: map_equality_iff)
lemma mapM_set:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>_. P\<rbrace>"
assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>_. Q x\<rbrace>"
assumes "\<And>x y. \<lbrakk> x \<in> set xs; y \<in> set xs \<rbrakk> \<Longrightarrow> \<lbrace>P and Q y\<rbrace> f x \<lbrace>\<lambda>_. Q y\<rbrace>"
shows "\<lbrace>P\<rbrace> mapM f xs \<lbrace>\<lambda>_ s. \<forall>x \<in> set xs. Q x s\<rbrace>"
using assms
proof (induct xs)
case Nil show ?case
by (simp add: mapM_def sequence_def) wp
next
case (Cons y ys)
have PQ_inv: "\<And>x. x \<in> set ys \<Longrightarrow> \<lbrace>P and Q y\<rbrace> f x \<lbrace>\<lambda>_. P and Q y\<rbrace>"
by (wpsimp wp: Cons)
show ?case
apply (simp add: mapM_Cons)
apply wp
apply (rule hoare_vcg_conj_lift)
apply (rule hoare_strengthen_post)
apply (rule mapM_wp')
apply (erule PQ_inv)
apply simp
apply (wp Cons|simp)+
done
qed
lemma mapM_set_inv:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>_. P\<rbrace>"
assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>_. Q x\<rbrace>"
assumes "\<And>x y. \<lbrakk> x \<in> set xs; y \<in> set xs \<rbrakk> \<Longrightarrow> \<lbrace>P and Q y\<rbrace> f x \<lbrace>\<lambda>_. Q y\<rbrace>"
shows "\<lbrace>P\<rbrace> mapM f xs \<lbrace>\<lambda>_ s. P s \<and> (\<forall>x \<in> set xs. Q x s)\<rbrace>"
apply (rule hoare_weaken_pre, rule hoare_vcg_conj_lift)
apply (rule mapM_wp', erule assms)
apply (rule mapM_set; rule assms; assumption)
apply simp
done
lemma mapM_x_wp:
assumes x: "\<And>x. x \<in> S \<Longrightarrow> \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>rv. P\<rbrace>"
shows "set xs \<subseteq> S \<Longrightarrow> \<lbrace>P\<rbrace> mapM_x f xs \<lbrace>\<lambda>rv. P\<rbrace>"
by (subst mapM_x_mapM) (wp mapM_wp x)
lemma no_fail_mapM':
assumes rl: "\<And>x. no_fail (\<lambda>_. P x) (f x)"
shows "no_fail (\<lambda>_. \<forall>x \<in> set xs. P x) (mapM f xs)"
proof (induct xs)
case Nil thus ?case by (simp add: mapM_def sequence_def)
next
case (Cons x xs)
have nf: "no_fail (\<lambda>_. P x) (f x)" by (rule rl)
have ih: "no_fail (\<lambda>_. \<forall>x \<in> set xs. P x) (mapM f xs)" by (rule Cons)
show ?case
apply (simp add: mapM_Cons)
apply (rule no_fail_pre)
apply (wp nf ih)
apply simp
done
qed
lemma det_mapM:
assumes x: "\<And>x. x \<in> S \<Longrightarrow> det (f x)"
shows "set xs \<subseteq> S \<Longrightarrow> det (mapM f xs)"
apply (induct xs)
apply (simp add: mapM_def sequence_def)
apply (simp add: mapM_Cons x)
done
lemma det_zipWithM_x:
assumes x: "\<And>x y. (x, y) \<in> set (zip xs ys) \<Longrightarrow> det (f x y)"
shows "det (zipWithM_x f xs ys)"
apply (simp add: zipWithM_x_mapM)
apply (rule bind_detI)
apply (rule det_mapM [where S="set (zip xs ys)"])
apply (clarsimp simp add: x)
apply simp
apply simp
done
lemma empty_fail_sequence_x :
assumes "\<And>m. m \<in> set ms \<Longrightarrow> empty_fail m"
shows "empty_fail (sequence_x ms)" using assms
by (induct ms) (auto simp: sequence_x_def)
lemma mapME_set:
assumes est: "\<And>x. \<lbrace>R\<rbrace> f x \<lbrace>P\<rbrace>, -"
and invp: "\<And>x y. \<lbrace>R and P x\<rbrace> f y \<lbrace>\<lambda>_. P x\<rbrace>, -"
and invr: "\<And>x. \<lbrace>R\<rbrace> f x \<lbrace>\<lambda>_. R\<rbrace>, -"
shows "\<lbrace>R\<rbrace> mapME f xs \<lbrace>\<lambda>rv s. \<forall>x \<in> set rv. P x s\<rbrace>, -"
proof (rule hoare_strengthen_postE_R [where Q' = "\<lambda>rv s. R s \<and> (\<forall>x \<in> set rv. P x s)"], induct xs)
case Nil
thus ?case by (simp add: mapME_Nil | wp returnOKE_R_wp)+
next
case (Cons y ys)
have minvp: "\<And>x. \<lbrace>R and P x\<rbrace> mapME f ys \<lbrace>\<lambda>_. P x\<rbrace>, -"
apply (rule hoare_pre)
apply (rule_tac Q' = "\<lambda>_ s. R s \<and> P x s" in hoare_strengthen_postE_R)
apply (wp mapME_wp' invr invp)+
apply simp
apply simp
apply simp
done
show ?case
apply (simp add: mapME_Cons)
apply (wp)
apply (rule_tac Q' = "\<lambda>xs s. (R s \<and> (\<forall>x \<in> set xs. P x s)) \<and> P rv s" in hoare_strengthen_postE_R)
apply (wp Cons.hyps minvp)
apply simp
apply (fold validE_R_def)
apply simp
apply (wp invr est)
apply simp
done
qed clarsimp
lemma empty_fail_mapM_x [simp]:
"(\<And>x. empty_fail (a x)) \<Longrightarrow> empty_fail (mapM_x a xs)"
apply (induct_tac xs)
apply (clarsimp simp: mapM_x_Nil)
apply (clarsimp simp: mapM_x_Cons)
done
lemma mapM_upd_inv:
assumes f: "\<And>x rv. (rv,s) \<in> fst (f x s) \<Longrightarrow> x \<in> set xs \<Longrightarrow> (rv, g s) \<in> fst (f x (g s))"
assumes inv: "\<And>x. \<lbrace>(=) s\<rbrace> f x \<lbrace>\<lambda>_. (=) s\<rbrace>"
shows "(rv,s) \<in> fst (mapM f xs s) \<Longrightarrow> (rv, g s) \<in> fst (mapM f xs (g s))"
using f inv
proof (induct xs arbitrary: rv s)
case Nil
thus ?case by (simp add: mapM_Nil return_def)
next
case (Cons z zs)
from Cons.prems
show ?case
apply (clarsimp simp: mapM_Cons in_monad)
apply (frule use_valid, assumption, rule refl)
apply clarsimp
apply (drule Cons.prems, simp)
apply (rule exI, erule conjI)
apply (drule Cons.hyps)
apply simp
apply assumption
apply simp
done
qed
lemma case_option_find_give_me_a_map:
"case_option a return (find f xs)
= liftM projl
(mapME (\<lambda>x. if (f x) then throwError x else returnOk ()) xs
>>=E (\<lambda>x. assert (\<forall>x \<in> set xs. \<not> f x)
>>= (\<lambda>_. liftM (Inl :: 'a \<Rightarrow> 'a + unit) a)))"
apply (induct xs)
apply simp
apply (simp add: liftM_def mapME_Nil)
apply (simp add: mapME_Cons split: if_split)
apply (clarsimp simp add: throwError_def bindE_def bind_assoc
liftM_def)
apply (rule bind_cong [OF refl])
apply (simp add: lift_def throwError_def returnOk_def split: sum.split)
done
end