-
Notifications
You must be signed in to change notification settings - Fork 1
/
List_Extra.thy
1326 lines (1084 loc) · 50.5 KB
/
List_Extra.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
708
709
710
711
712
713
714
715
716
717
718
719
720
721
722
723
724
725
726
727
728
729
730
731
732
733
734
735
736
737
738
739
740
741
742
743
744
745
746
747
748
749
750
751
752
753
754
755
756
757
758
759
760
761
762
763
764
765
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
793
794
795
796
797
798
799
800
801
802
803
804
805
806
807
808
809
810
811
812
813
814
815
816
817
818
819
820
821
822
823
824
825
826
827
828
829
830
831
832
833
834
835
836
837
838
839
840
841
842
843
844
845
846
847
848
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
874
875
876
877
878
879
880
881
882
883
884
885
886
887
888
889
890
891
892
893
894
895
896
897
898
899
900
901
902
903
904
905
906
907
908
909
910
911
912
913
914
915
916
917
918
919
920
921
922
923
924
925
926
927
928
929
930
931
932
933
934
935
936
937
938
939
940
941
942
943
944
945
946
947
948
949
950
951
952
953
954
955
956
957
958
959
960
961
962
963
964
965
966
967
968
969
970
971
972
973
974
975
976
977
978
979
980
981
982
983
984
985
986
987
988
989
990
991
992
993
994
995
996
997
998
999
1000
(*****************************************************************************************)
(* Project: Isabelle/UTP Toolkit *)
(* File: List_Extra.thy *)
(* Authors: Simon Foster, Pedro Ribeiro, and Frank Zeyda *)
(* Emails: simon.foster@york.ac.uk, pedro.ribeiro@york.ac.uk, and frank.zeyda@york.ac.uk *)
(*****************************************************************************************)
section \<open> Lists: extra functions and properties \<close>
theory List_Extra
imports
"HOL-Library.Sublist"
"HOL-Library.Monad_Syntax"
"HOL-Library.Prefix_Order"
"Optics.Lens_Instances"
begin
subsection \<open> Useful Abbreviations \<close>
abbreviation "list_sum xs \<equiv> foldr (+) xs 0"
subsection \<open> Sets \<close>
lemma set_Fpow [simp]: "set xs \<in> Fpow A \<longleftrightarrow> set xs \<subseteq> A"
by (auto simp add: Fpow_def)
lemma Fpow_as_Pow: "finite A \<Longrightarrow> Fpow A = Pow A"
by (auto simp add: Pow_def Fpow_def finite_subset)
lemma Fpow_set [code]:
"Fpow (set []) = {{}}"
"Fpow (set (x # xs)) = (let A = Fpow (set xs) in A \<union> insert x ` A)"
by (simp_all add: Fpow_as_Pow Pow_set del: set_simps)
subsection \<open> Folds \<close>
context abel_semigroup
begin
lemma foldr_snoc: "foldr (\<^bold>*) (xs @ [x]) k = (foldr (\<^bold>*) xs k) \<^bold>* x"
by (induct xs, simp_all add: commute left_commute)
end
subsection \<open> List Lookup \<close>
text \<open>
The following variant of the standard @{text nth} function returns @{text "\<bottom>"} if the index is
out of range.
\<close>
primrec
nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>\<^sub>l" [90, 0] 91)
where
"[]\<langle>i\<rangle>\<^sub>l = None"
| "(x # xs)\<langle>i\<rangle>\<^sub>l = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>\<^sub>l)"
lemma nth_el_appendl[simp]: "i < length xs \<Longrightarrow> (xs @ ys)\<langle>i\<rangle>\<^sub>l = xs\<langle>i\<rangle>\<^sub>l"
apply (induct xs arbitrary: i)
apply simp
apply (case_tac i)
apply simp_all
done
lemma nth_el_appendr[simp]: "length xs \<le> i \<Longrightarrow> (xs @ ys)\<langle>i\<rangle>\<^sub>l = ys\<langle>i - length xs\<rangle>\<^sub>l"
apply (induct xs arbitrary: i)
apply simp
apply (case_tac i)
apply simp_all
done
subsection \<open> Extra List Theorems \<close>
subsubsection \<open> Map \<close>
lemma map_nth_Cons_atLeastLessThan:
"map (nth (x # xs)) [Suc m..<n] = map (nth xs) [m..<n - 1]"
proof -
have "nth xs = nth (x # xs) \<circ> Suc"
by auto
hence "map (nth xs) [m..<n - 1] = map (nth (x # xs) \<circ> Suc) [m..<n - 1]"
by simp
also have "... = map (nth (x # xs)) (map Suc [m..<n - 1])"
by simp
also have "... = map (nth (x # xs)) [Suc m..<n]"
by (metis Suc_diff_1 le_0_eq length_upt list.simps(8) list.size(3) map_Suc_upt not_less upt_0)
finally show ?thesis ..
qed
subsubsection \<open> Sorted Lists \<close>
lemma sorted_last [simp]: "\<lbrakk> x \<in> set xs; sorted xs \<rbrakk> \<Longrightarrow> x \<le> last xs"
by (induct xs, auto)
lemma sorted_prefix:
assumes "xs \<le> ys" "sorted ys"
shows "sorted xs"
proof -
obtain zs where "ys = xs @ zs"
using Prefix_Order.prefixE assms(1) by auto
thus ?thesis
using assms(2) sorted_append by blast
qed
lemma sorted_map: "\<lbrakk> sorted xs; mono f \<rbrakk> \<Longrightarrow> sorted (map f xs)"
by (simp add: monoD sorted_iff_nth_mono)
lemma sorted_distinct [intro]: "\<lbrakk> sorted (xs); distinct(xs) \<rbrakk> \<Longrightarrow> (\<forall> i<length xs - 1. xs!i < xs!(i + 1))"
apply (induct xs)
apply (auto)
apply (metis (no_types, opaque_lifting) Suc_leI Suc_less_eq Suc_pred gr0_conv_Suc not_le not_less_iff_gr_or_eq nth_Cons_Suc nth_mem nth_non_equal_first_eq)
done
text \<open> The concatenation of two lists is sorted provided (1) both the lists are sorted, and (2)
the final and first elements are ordered. \<close>
lemma sorted_append_middle:
"sorted(xs@ys) = (sorted xs \<and> sorted ys \<and> (xs \<noteq> [] \<and> ys \<noteq> [] \<longrightarrow> xs!(length xs - 1) \<le> ys!0))"
proof -
have "\<And>x y. \<lbrakk> sorted xs; sorted ys; xs ! (length xs - Suc 0) \<le> ys ! 0 \<rbrakk> \<Longrightarrow> x \<in> set xs \<Longrightarrow> y \<in> set ys \<Longrightarrow> x \<le> y"
proof -
fix x y
assume "sorted xs" "sorted ys" "xs ! (length xs - Suc 0) \<le> ys ! 0" "x \<in> set xs" "y \<in> set ys"
moreover then obtain i j where i: "x = xs!i" "i < length xs" and j: "y = ys!j" "j < length ys"
by (auto simp add: in_set_conv_nth)
moreover have "xs ! i \<le> xs!(length xs - 1)"
by (metis One_nat_def Suc_diff_Suc Suc_leI Suc_le_mono \<open>i < length xs\<close> \<open>sorted xs\<close> diff_less diff_zero gr_implies_not_zero nat.simps(3) sorted_iff_nth_mono zero_less_iff_neq_zero)
moreover have "ys!0 \<le> ys ! j"
by (simp add: calculation(2) calculation(9) sorted_nth_mono)
ultimately have "xs ! i \<le> ys ! j"
by (metis One_nat_def dual_order.trans)
thus "x \<le> y"
by (simp add: i j)
qed
thus ?thesis
by (auto simp add: sorted_append)
qed
text \<open> Is the given list a permutation of the given set? \<close>
definition is_sorted_list_of_set :: "('a::ord) set \<Rightarrow> 'a list \<Rightarrow> bool" where
"is_sorted_list_of_set A xs = ((\<forall> i<length(xs) - 1. xs!i < xs!(i + 1)) \<and> set(xs) = A)"
lemma sorted_is_sorted_list_of_set:
assumes "is_sorted_list_of_set A xs"
shows "sorted(xs)" and "distinct(xs)"
using assms proof (induct xs arbitrary: A)
show "sorted []"
by auto
next
show "distinct []"
by auto
next
fix A :: "'a set"
case (Cons x xs) note hyps = this
assume isl: "is_sorted_list_of_set A (x # xs)"
hence srt: "(\<forall>i<length xs - Suc 0. xs ! i < xs ! Suc i)"
using less_diff_conv by (auto simp add: is_sorted_list_of_set_def)
with hyps(1) have srtd: "sorted xs"
by (simp add: is_sorted_list_of_set_def)
with isl show "sorted (x # xs)"
apply (auto simp add: is_sorted_list_of_set_def)
apply (metis (mono_tags, lifting) all_nth_imp_all_set less_le_trans linorder_not_less not_less_iff_gr_or_eq nth_Cons_0 sorted_iff_nth_mono zero_order(3))
done
from srt hyps(2) have "distinct xs"
by (simp add: is_sorted_list_of_set_def)
with isl show "distinct (x # xs)"
proof -
have "(\<forall>n. \<not> n < length (x # xs) - 1 \<or> (x # xs) ! n < (x # xs) ! (n + 1)) \<and> set (x # xs) = A"
by (meson \<open>is_sorted_list_of_set A (x # xs)\<close> is_sorted_list_of_set_def)
then show ?thesis
by (metis \<open>distinct xs\<close> add.commute add_diff_cancel_left' distinct.simps(2) leD length_Cons length_greater_0_conv length_pos_if_in_set less_le nth_Cons_0 nth_Cons_Suc plus_1_eq_Suc set_ConsD sorted_wrt.elims(2) srtd)
qed
qed
term sorted
lemma is_sorted_list_of_set_alt_def:
"is_sorted_list_of_set A xs \<longleftrightarrow> sorted (xs) \<and> distinct (xs) \<and> set(xs) = A"
apply (auto intro: sorted_is_sorted_list_of_set)
apply (auto simp add: is_sorted_list_of_set_def)
apply (metis Nat.add_0_right One_nat_def add_Suc_right sorted_distinct)
done
definition sorted_list_of_set_alt :: "('a::ord) set \<Rightarrow> 'a list" where
"sorted_list_of_set_alt A =
(if (A = {}) then [] else (THE xs. is_sorted_list_of_set A xs))"
lemma is_sorted_list_of_set:
"finite A \<Longrightarrow> is_sorted_list_of_set A (sorted_list_of_set A)"
using sorted_distinct sorted_list_of_set(2) by (fastforce simp add: is_sorted_list_of_set_def)
lemma sorted_list_of_set_other_def:
"finite A \<Longrightarrow> sorted_list_of_set(A) = (THE xs. sorted(xs) \<and> distinct(xs) \<and> set xs = A)"
apply (rule sym)
apply (rule the_equality)
apply (auto)
apply (simp add: sorted_distinct_set_unique)
done
lemma sorted_list_of_set_alt [simp]:
"finite A \<Longrightarrow> sorted_list_of_set_alt(A) = sorted_list_of_set(A)"
apply (rule sym)
apply (auto simp add: sorted_list_of_set_alt_def is_sorted_list_of_set_alt_def sorted_list_of_set_other_def)
done
text \<open> Sorting lists according to a relation \<close>
definition is_sorted_list_of_set_by :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" where
"is_sorted_list_of_set_by R A xs = ((\<forall> i<length(xs) - 1. (xs!i, xs!(i + 1)) \<in> R) \<and> set(xs) = A)"
definition sorted_list_of_set_by :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a list" where
"sorted_list_of_set_by R A = (THE xs. is_sorted_list_of_set_by R A xs)"
definition fin_set_lexord :: "'a rel \<Rightarrow> 'a set rel" where
"fin_set_lexord R = {(A, B). finite A \<and> finite B \<and>
(\<exists> xs ys. is_sorted_list_of_set_by R A xs \<and> is_sorted_list_of_set_by R B ys
\<and> (xs, ys) \<in> lexord R)}"
lemma is_sorted_list_of_set_by_mono:
"\<lbrakk> R \<subseteq> S; is_sorted_list_of_set_by R A xs \<rbrakk> \<Longrightarrow> is_sorted_list_of_set_by S A xs"
by (auto simp add: is_sorted_list_of_set_by_def)
lemma lexord_mono':
"\<lbrakk> (\<And> x y. f x y \<longrightarrow> g x y); (xs, ys) \<in> lexord {(x, y). f x y} \<rbrakk> \<Longrightarrow> (xs, ys) \<in> lexord {(x, y). g x y}"
by (metis case_prodD case_prodI lexord_take_index_conv mem_Collect_eq)
lemma fin_set_lexord_mono [mono]:
"(\<And> x y. f x y \<longrightarrow> g x y) \<Longrightarrow> (xs, ys) \<in> fin_set_lexord {(x, y). f x y} \<longrightarrow> (xs, ys) \<in> fin_set_lexord {(x, y). g x y}"
proof
assume
fin: "(xs, ys) \<in> fin_set_lexord {(x, y). f x y}" and
hyp: "(\<And> x y. f x y \<longrightarrow> g x y)"
from fin have "finite xs" "finite ys"
using fin_set_lexord_def by fastforce+
with fin hyp show "(xs, ys) \<in> fin_set_lexord {(x, y). g x y}"
apply (auto simp add: fin_set_lexord_def)
apply (rename_tac xs' ys')
apply (rule_tac x="xs'" in exI)
apply (auto)
apply (metis case_prodD case_prodI is_sorted_list_of_set_by_def mem_Collect_eq)
apply (metis case_prodD case_prodI is_sorted_list_of_set_by_def lexord_mono' mem_Collect_eq)
done
qed
definition distincts :: "'a set \<Rightarrow> 'a list set" where
"distincts A = {xs \<in> lists A. distinct(xs)}"
lemma tl_element:
"\<lbrakk> x \<in> set xs; x \<noteq> hd(xs) \<rbrakk> \<Longrightarrow> x \<in> set(tl(xs))"
by (metis in_set_insert insert_Nil list.collapse list.distinct(2) set_ConsD)
subsubsection \<open> List Update \<close>
lemma listsum_update:
fixes xs :: "'a::ring list"
assumes "i < length xs"
shows "list_sum (xs[i := v]) = list_sum xs - xs ! i + v"
using assms proof (induct xs arbitrary: i)
case Nil
then show ?case by (simp)
next
case (Cons a xs)
then show ?case
proof (cases i)
case 0
thus ?thesis
by (simp add: add.commute)
next
case (Suc i')
with Cons show ?thesis
by (auto)
qed
qed
subsubsection \<open> Drop While and Take While \<close>
lemma dropWhile_sorted_le_above:
"\<lbrakk> sorted xs; x \<in> set (dropWhile (\<lambda> x. x \<le> n) xs) \<rbrakk> \<Longrightarrow> x > n"
apply (induct xs)
apply (auto)
apply (rename_tac a xs)
apply (case_tac "a \<le> n")
apply (auto)
done
lemma set_dropWhile_le:
"sorted xs \<Longrightarrow> set (dropWhile (\<lambda> x. x \<le> n) xs) = {x\<in>set xs. x > n}"
apply (induct xs)
apply (simp)
apply (rename_tac x xs)
apply (subgoal_tac "sorted xs")
apply (simp)
apply (safe)
apply (auto)
done
lemma set_takeWhile_less_sorted:
"\<lbrakk> sorted I; x \<in> set I; x < n \<rbrakk> \<Longrightarrow> x \<in> set (takeWhile (\<lambda>x. x < n) I)"
proof (induct I arbitrary: x)
case Nil thus ?case
by (simp)
next
case (Cons a I) thus ?case
by auto
qed
lemma nth_le_takeWhile_ord: "\<lbrakk> sorted xs; i \<ge> length (takeWhile (\<lambda> x. x \<le> n) xs); i < length xs \<rbrakk> \<Longrightarrow> n \<le> xs ! i"
apply (induct xs arbitrary: i, auto)
apply (rename_tac x xs i)
apply (case_tac "x \<le> n")
apply (auto)
apply (metis One_nat_def Suc_eq_plus1 le_less_linear le_less_trans less_imp_le list.size(4) nth_mem set_ConsD)
done
lemma length_takeWhile_less:
"\<lbrakk> a \<in> set xs; \<not> P a \<rbrakk> \<Longrightarrow> length (takeWhile P xs) < length xs"
by (metis in_set_conv_nth length_takeWhile_le nat_neq_iff not_less set_takeWhileD takeWhile_nth)
lemma nth_length_takeWhile_less:
"\<lbrakk> sorted xs; distinct xs; (\<exists> a \<in> set xs. a \<ge> n) \<rbrakk> \<Longrightarrow> xs ! length (takeWhile (\<lambda>x. x < n) xs) \<ge> n"
by (induct xs, auto)
subsubsection \<open> Last and But Last \<close>
lemma length_gt_zero_butlast_concat:
assumes "length ys > 0"
shows "butlast (xs @ ys) = xs @ (butlast ys)"
using assms by (metis butlast_append length_greater_0_conv)
lemma length_eq_zero_butlast_concat:
assumes "length ys = 0"
shows "butlast (xs @ ys) = butlast xs"
using assms by (metis append_Nil2 length_0_conv)
lemma butlast_single_element:
shows "butlast [e] = []"
by (metis butlast.simps(2))
lemma last_single_element:
shows "last [e] = e"
by (metis last.simps)
lemma length_zero_last_concat:
assumes "length t = 0"
shows "last (s @ t) = last s"
by (metis append_Nil2 assms length_0_conv)
lemma length_gt_zero_last_concat:
assumes "length t > 0"
shows "last (s @ t) = last t"
by (metis assms last_append length_greater_0_conv)
subsubsection \<open> Prefixes and Strict Prefixes \<close>
lemma prefix_length_eq:
"\<lbrakk> length xs = length ys; prefix xs ys \<rbrakk> \<Longrightarrow> xs = ys"
by (metis not_equal_is_parallel parallel_def)
lemma prefix_Cons_elim [elim]:
assumes "prefix (x # xs) ys"
obtains ys' where "ys = x # ys'" "prefix xs ys'"
using assms
by (metis append_Cons prefix_def)
lemma prefix_map_inj:
"\<lbrakk> inj_on f (set xs \<union> set ys); prefix (map f xs) (map f ys) \<rbrakk> \<Longrightarrow>
prefix xs ys"
apply (induct xs arbitrary:ys)
apply (simp_all)
apply (erule prefix_Cons_elim)
apply (auto)
apply (metis image_insert insertI1 insert_Diff_if singletonE)
done
lemma prefix_map_inj_eq [simp]:
"inj_on f (set xs \<union> set ys) \<Longrightarrow>
prefix (map f xs) (map f ys) \<longleftrightarrow> prefix xs ys"
using map_mono_prefix prefix_map_inj by blast
lemma strict_prefix_Cons_elim [elim]:
assumes "strict_prefix (x # xs) ys"
obtains ys' where "ys = x # ys'" "strict_prefix xs ys'"
using assms
by (metis Sublist.strict_prefixE' Sublist.strict_prefixI' append_Cons)
lemma strict_prefix_map_inj:
"\<lbrakk> inj_on f (set xs \<union> set ys); strict_prefix (map f xs) (map f ys) \<rbrakk> \<Longrightarrow>
strict_prefix xs ys"
apply (induct xs arbitrary:ys)
apply (auto)
using prefix_bot.not_eq_extremum apply fastforce
apply (erule strict_prefix_Cons_elim)
apply (auto)
apply (metis (opaque_lifting, full_types) image_insert insertI1 insert_Diff_if singletonE)
done
lemma strict_prefix_map_inj_eq [simp]:
"inj_on f (set xs \<union> set ys) \<Longrightarrow>
strict_prefix (map f xs) (map f ys) \<longleftrightarrow> strict_prefix xs ys"
by (simp add: inj_on_map_eq_map strict_prefix_def)
lemma prefix_drop:
"\<lbrakk> drop (length xs) ys = zs; prefix xs ys \<rbrakk>
\<Longrightarrow> ys = xs @ zs"
by (metis append_eq_conv_conj prefix_def)
lemma list_append_prefixD [dest]: "x @ y \<le> z \<Longrightarrow> x \<le> z"
using append_prefixD less_eq_list_def by blast
lemma prefix_not_empty:
assumes "strict_prefix xs ys" and "xs \<noteq> []"
shows "ys \<noteq> []"
using Sublist.strict_prefix_simps(1) assms(1) by blast
lemma prefix_not_empty_length_gt_zero:
assumes "strict_prefix xs ys" and "xs \<noteq> []"
shows "length ys > 0"
using assms prefix_not_empty by auto
lemma butlast_prefix_suffix_not_empty:
assumes "strict_prefix (butlast xs) ys"
shows "ys \<noteq> []"
using assms prefix_not_empty_length_gt_zero by fastforce
lemma prefix_and_concat_prefix_is_concat_prefix:
assumes "prefix s t" "prefix (e @ t) u"
shows "prefix (e @ s) u"
using Sublist.same_prefix_prefix assms(1) assms(2) prefix_order.dual_order.trans by blast
lemma prefix_eq_exists:
"prefix s t \<longleftrightarrow> (\<exists>xs . s @ xs = t)"
using prefix_def by auto
lemma strict_prefix_eq_exists:
"strict_prefix s t \<longleftrightarrow> (\<exists>xs . s @ xs = t \<and> (length xs) > 0)"
using prefix_def strict_prefix_def by auto
lemma butlast_strict_prefix_eq_butlast:
assumes "length s = length t" and "strict_prefix (butlast s) t"
shows "strict_prefix (butlast s) t \<longleftrightarrow> (butlast s) = (butlast t)"
by (metis append_butlast_last_id append_eq_append_conv assms(1) assms(2) length_0_conv length_butlast strict_prefix_eq_exists)
lemma butlast_eq_if_eq_length_and_prefix:
assumes "length s > 0" "length z > 0"
"length s = length z" "strict_prefix (butlast s) t" "strict_prefix (butlast z) t"
shows "(butlast s) = (butlast z)"
using assms by (auto simp add:strict_prefix_eq_exists)
lemma prefix_imp_length_lteq:
assumes "prefix s t"
shows "length s \<le> length t"
using assms by (simp add: Sublist.prefix_length_le)
lemma prefix_imp_length_not_gt:
assumes "prefix s t"
shows "\<not> length t < length s"
using assms by (simp add: Sublist.prefix_length_le leD)
lemma prefix_and_eq_length_imp_eq_list:
assumes "prefix s t" and "length t = length s"
shows "s=t"
using assms by (simp add: prefix_length_eq)
lemma butlast_prefix_imp_length_not_gt:
assumes "length s > 0" "strict_prefix (butlast s) t"
shows "\<not> (length t < length s)"
using assms prefix_length_less by fastforce
lemma length_not_gt_iff_eq_length:
assumes "length s > 0" and "strict_prefix (butlast s) t"
shows "(\<not> (length s < length t)) = (length s = length t)"
proof -
have "(\<not> (length s < length t)) = ((length t < length s) \<or> (length s = length t))"
by (metis not_less_iff_gr_or_eq)
also have "... = (length s = length t)"
using assms
by (simp add:butlast_prefix_imp_length_not_gt)
finally show ?thesis .
qed
lemma list_prefix_iff:
"(prefix xs ys \<longleftrightarrow> (length xs \<le> length ys \<and> (\<forall> i<length xs. xs!i = ys!i)))"
apply (auto)
apply (simp add: prefix_imp_length_lteq)
apply (metis nth_append prefix_def)
apply (metis nth_take_lemma order_refl take_all take_is_prefix)
done
lemma list_le_prefix_iff:
"(xs \<le> ys \<longleftrightarrow> (length xs \<le> length ys \<and> (\<forall> i<length xs. xs!i = ys!i)))"
by (simp add: less_eq_list_def list_prefix_iff)
text \<open> Greatest common prefix \<close>
fun gcp :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"gcp [] ys = []" |
"gcp (x # xs) (y # ys) = (if (x = y) then x # gcp xs ys else [])" |
"gcp _ _ = []"
lemma gcp_right [simp]: "gcp xs [] = []"
by (induct xs, auto)
lemma gcp_append [simp]: "gcp (xs @ ys) (xs @ zs) = xs @ gcp ys zs"
by (induct xs, auto)
lemma gcp_lb1: "prefix (gcp xs ys) xs"
apply (induct xs arbitrary: ys, auto)
apply (case_tac ys, auto)
done
lemma gcp_lb2: "prefix (gcp xs ys) ys"
apply (induct ys arbitrary: xs, auto)
apply (case_tac xs, auto)
done
interpretation prefix_semilattice: semilattice_inf gcp prefix strict_prefix
proof
fix xs ys :: "'a list"
show "prefix (gcp xs ys) xs"
by (induct xs arbitrary: ys, auto, case_tac ys, auto)
show "prefix (gcp xs ys) ys"
by (induct ys arbitrary: xs, auto, case_tac xs, auto)
next
fix xs ys zs :: "'a list"
assume "prefix xs ys" "prefix xs zs"
thus "prefix xs (gcp ys zs)"
by (simp add: prefix_def, auto)
qed
subsubsection \<open> Lexicographic Order \<close>
lemma lexord_append:
assumes "(xs\<^sub>1 @ ys\<^sub>1, xs\<^sub>2 @ ys\<^sub>2) \<in> lexord R" "length(xs\<^sub>1) = length(xs\<^sub>2)"
shows "(xs\<^sub>1, xs\<^sub>2) \<in> lexord R \<or> (xs\<^sub>1 = xs\<^sub>2 \<and> (ys\<^sub>1, ys\<^sub>2) \<in> lexord R)"
using assms
proof (induct xs\<^sub>2 arbitrary: xs\<^sub>1)
case (Cons x\<^sub>2 xs\<^sub>2') note hyps = this
from hyps(3) obtain x\<^sub>1 xs\<^sub>1' where xs\<^sub>1: "xs\<^sub>1 = x\<^sub>1 # xs\<^sub>1'" "length(xs\<^sub>1') = length(xs\<^sub>2')"
by (auto, metis Suc_length_conv)
with hyps(2) have xcases: "(x\<^sub>1, x\<^sub>2) \<in> R \<or> (xs\<^sub>1' @ ys\<^sub>1, xs\<^sub>2' @ ys\<^sub>2) \<in> lexord R"
by (auto)
show ?case
proof (cases "(x\<^sub>1, x\<^sub>2) \<in> R")
case True with xs\<^sub>1 show ?thesis
by (auto)
next
case False
with xcases have "(xs\<^sub>1' @ ys\<^sub>1, xs\<^sub>2' @ ys\<^sub>2) \<in> lexord R"
by (auto)
with hyps(1) xs\<^sub>1 have dichot: "(xs\<^sub>1', xs\<^sub>2') \<in> lexord R \<or> (xs\<^sub>1' = xs\<^sub>2' \<and> (ys\<^sub>1, ys\<^sub>2) \<in> lexord R)"
by (auto)
have "x\<^sub>1 = x\<^sub>2"
using False hyps(2) xs\<^sub>1(1) by auto
with dichot xs\<^sub>1 show ?thesis
by (simp)
qed
next
case Nil thus ?case
by auto
qed
lemma strict_prefix_lexord_rel:
"strict_prefix xs ys \<Longrightarrow> (xs, ys) \<in> lexord R"
by (metis Sublist.strict_prefixE' lexord_append_rightI)
lemma strict_prefix_lexord_left:
assumes "trans R" "(xs, ys) \<in> lexord R" "strict_prefix xs' xs"
shows "(xs', ys) \<in> lexord R"
by (metis assms lexord_trans strict_prefix_lexord_rel)
lemma prefix_lexord_right:
assumes "trans R" "(xs, ys) \<in> lexord R" "strict_prefix ys ys'"
shows "(xs, ys') \<in> lexord R"
by (metis assms lexord_trans strict_prefix_lexord_rel)
lemma lexord_eq_length:
assumes "(xs, ys) \<in> lexord R" "length xs = length ys"
shows "\<exists> i. (xs!i, ys!i) \<in> R \<and> i < length xs \<and> (\<forall> j<i. xs!j = ys!j)"
using assms proof (induct xs arbitrary: ys)
case (Cons x xs) note hyps = this
then obtain y ys' where ys: "ys = y # ys'" "length ys' = length xs"
by (metis Suc_length_conv)
show ?case
proof (cases "(x, y) \<in> R")
case True with ys show ?thesis
by (rule_tac x="0" in exI, simp)
next
case False
with ys hyps(2) have xy: "x = y" "(xs, ys') \<in> lexord R"
by auto
with hyps(1,3) ys obtain i where "(xs!i, ys'!i) \<in> R" "i < length xs" "(\<forall> j<i. xs!j = ys'!j)"
by force
with xy ys show ?thesis
apply (rule_tac x="Suc i" in exI)
apply (auto simp add: less_Suc_eq_0_disj)
done
qed
next
case Nil thus ?case by (auto)
qed
lemma lexord_intro_elems:
assumes "length xs > i" "length ys > i" "(xs!i, ys!i) \<in> R" "\<forall> j<i. xs!j = ys!j"
shows "(xs, ys) \<in> lexord R"
using assms proof (induct i arbitrary: xs ys)
case 0 thus ?case
by (auto, metis lexord_cons_cons list.exhaust nth_Cons_0)
next
case (Suc i) note hyps = this
then obtain x' y' xs' ys' where "xs = x' # xs'" "ys = y' # ys'"
by (metis Suc_length_conv Suc_lessE)
moreover with hyps(5) have "\<forall>j<i. xs' ! j = ys' ! j"
by (auto)
ultimately show ?case using hyps
by (auto)
qed
subsection \<open> Distributed Concatenation \<close>
definition uncurry :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b \<Rightarrow> 'c)" where
[simp]: "uncurry f = (\<lambda>(x, y). f x y)"
definition dist_concat ::
"'a list set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" (infixr "\<^sup>\<frown>" 100) where
"dist_concat ls1 ls2 = (uncurry (@) ` (ls1 \<times> ls2))"
lemma dist_concat_left_empty [simp]:
"{} \<^sup>\<frown> ys = {}"
by (simp add: dist_concat_def)
lemma dist_concat_right_empty [simp]:
"xs \<^sup>\<frown> {} = {}"
by (simp add: dist_concat_def)
lemma dist_concat_insert [simp]:
"insert l ls1 \<^sup>\<frown> ls2 = ((@) l ` ( ls2)) \<union> (ls1 \<^sup>\<frown> ls2)"
by (auto simp add: dist_concat_def)
subsection \<open> List Domain and Range \<close>
abbreviation seq_dom :: "'a list \<Rightarrow> nat set" ("dom\<^sub>l") where
"seq_dom xs \<equiv> {0..<length xs}"
abbreviation (input) seq_ran :: "'a list \<Rightarrow> 'a set" ("ran\<^sub>l") where
"seq_ran xs \<equiv> set xs"
subsection \<open> Extracting List Elements \<close>
definition seq_extract :: "nat set \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<upharpoonleft>\<^sub>l" 80) where
"seq_extract A xs = nths xs A"
lemma seq_extract_Nil [simp]: "A \<upharpoonleft>\<^sub>l [] = []"
by (simp add: seq_extract_def)
lemma seq_extract_Cons:
"A \<upharpoonleft>\<^sub>l (x # xs) = (if 0 \<in> A then [x] else []) @ {j. Suc j \<in> A} \<upharpoonleft>\<^sub>l xs"
by (simp add: seq_extract_def nths_Cons)
lemma seq_extract_empty [simp]: "{} \<upharpoonleft>\<^sub>l xs = []"
by (simp add: seq_extract_def)
lemma seq_extract_ident [simp]: "{0..<length xs} \<upharpoonleft>\<^sub>l xs = xs"
unfolding list_eq_iff_nth_eq
by (auto simp add: seq_extract_def length_nths atLeast0LessThan)
lemma seq_extract_split:
assumes "i \<le> length xs"
shows "{0..<i} \<upharpoonleft>\<^sub>l xs @ {i..<length xs} \<upharpoonleft>\<^sub>l xs = xs"
using assms
proof (induct xs arbitrary: i)
case Nil thus ?case by (simp add: seq_extract_def)
next
case (Cons x xs) note hyp = this
have "{j. Suc j < i} = {0..<i - 1}"
by (auto)
moreover have "{j. i \<le> Suc j \<and> j < length xs} = {i - 1..<length xs}"
by (auto)
ultimately show ?case
using hyp by (force simp add: seq_extract_def nths_Cons)
qed
lemma seq_extract_append:
"A \<upharpoonleft>\<^sub>l (xs @ ys) = (A \<upharpoonleft>\<^sub>l xs) @ ({j. j + length xs \<in> A} \<upharpoonleft>\<^sub>l ys)"
by (simp add: seq_extract_def nths_append)
lemma seq_extract_range: "A \<upharpoonleft>\<^sub>l xs = (A \<inter> dom\<^sub>l(xs)) \<upharpoonleft>\<^sub>l xs"
apply (auto simp add: seq_extract_def nths_def)
apply (metis (no_types, lifting) atLeastLessThan_iff filter_cong in_set_zip nth_mem set_upt)
done
lemma seq_extract_out_of_range:
"A \<inter> dom\<^sub>l(xs) = {} \<Longrightarrow> A \<upharpoonleft>\<^sub>l xs = []"
by (metis seq_extract_def seq_extract_range nths_empty)
lemma seq_extract_length [simp]:
"length (A \<upharpoonleft>\<^sub>l xs) = card (A \<inter> dom\<^sub>l(xs))"
proof -
have "{i. i < length(xs) \<and> i \<in> A} = (A \<inter> {0..<length(xs)})"
by (auto)
thus ?thesis
by (simp add: seq_extract_def length_nths)
qed
lemma seq_extract_Cons_atLeastLessThan:
assumes "m < n"
shows "{m..<n} \<upharpoonleft>\<^sub>l (x # xs) = (if (m = 0) then x # ({0..<n-1} \<upharpoonleft>\<^sub>l xs) else {m-1..<n-1} \<upharpoonleft>\<^sub>l xs)"
proof -
have "{j. Suc j < n} = {0..<n - Suc 0}"
by (auto)
moreover have "{j. m \<le> Suc j \<and> Suc j < n} = {m - Suc 0..<n - Suc 0}"
by (auto)
ultimately show ?thesis using assms
by (auto simp add: seq_extract_Cons)
qed
lemma seq_extract_singleton:
assumes "i < length xs"
shows "{i} \<upharpoonleft>\<^sub>l xs = [xs ! i]"
using assms
apply (induct xs arbitrary: i)
apply (auto simp add: seq_extract_Cons)
apply (rename_tac xs i)
apply (subgoal_tac "{j. Suc j = i} = {i - 1}")
apply (auto)
done
lemma seq_extract_as_map:
assumes "m < n" "n \<le> length xs"
shows "{m..<n} \<upharpoonleft>\<^sub>l xs = map (nth xs) [m..<n]"
using assms proof (induct xs arbitrary: m n)
case Nil thus ?case by simp
next
case (Cons x xs)
have "[m..<n] = m # [m+1..<n]"
using Cons.prems(1) upt_eq_Cons_conv by blast
moreover have "map (nth (x # xs)) [Suc m..<n] = map (nth xs) [m..<n-1]"
by (simp add: map_nth_Cons_atLeastLessThan)
ultimately show ?case
using Cons upt_rec
by (auto simp add: seq_extract_Cons_atLeastLessThan)
qed
lemma seq_append_as_extract:
"xs = ys @ zs \<longleftrightarrow> (\<exists> i\<le>length(xs). ys = {0..<i} \<upharpoonleft>\<^sub>l xs \<and> zs = {i..<length(xs)} \<upharpoonleft>\<^sub>l xs)"
proof
assume xs: "xs = ys @ zs"
moreover have "ys = {0..<length ys} \<upharpoonleft>\<^sub>l (ys @ zs)"
by (simp add: seq_extract_append)
moreover have "zs = {length ys..<length ys + length zs} \<upharpoonleft>\<^sub>l (ys @ zs)"
proof -
have "{length ys..<length ys + length zs} \<inter> {0..<length ys} = {}"
by auto
moreover have s1: "{j. j < length zs} = {0..<length zs}"
by auto
ultimately show ?thesis
by (simp add: seq_extract_append seq_extract_out_of_range)
qed
ultimately show "(\<exists> i\<le>length(xs). ys = {0..<i} \<upharpoonleft>\<^sub>l xs \<and> zs = {i..<length(xs)} \<upharpoonleft>\<^sub>l xs)"
by (rule_tac x="length ys" in exI, auto)
next
assume "\<exists>i\<le>length xs. ys = {0..<i} \<upharpoonleft>\<^sub>l xs \<and> zs = {i..<length xs} \<upharpoonleft>\<^sub>l xs"
thus "xs = ys @ zs"
by (auto simp add: seq_extract_split)
qed
subsection \<open> Filtering a list according to a set \<close>
definition seq_filter :: "'a list \<Rightarrow> 'a set \<Rightarrow> 'a list" (infix "\<restriction>\<^sub>l" 80) where
"seq_filter xs A = filter (\<lambda> x. x \<in> A) xs"
lemma seq_filter_Cons_in [simp]:
"x \<in> cs \<Longrightarrow> (x # xs) \<restriction>\<^sub>l cs = x # (xs \<restriction>\<^sub>l cs)"
by (simp add: seq_filter_def)
lemma seq_filter_Cons_out [simp]:
"x \<notin> cs \<Longrightarrow> (x # xs) \<restriction>\<^sub>l cs = (xs \<restriction>\<^sub>l cs)"
by (simp add: seq_filter_def)
lemma seq_filter_Nil [simp]: "[] \<restriction>\<^sub>l A = []"
by (simp add: seq_filter_def)
lemma seq_filter_empty [simp]: "xs \<restriction>\<^sub>l {} = []"
by (simp add: seq_filter_def)
lemma seq_filter_append: "(xs @ ys) \<restriction>\<^sub>l A = (xs \<restriction>\<^sub>l A) @ (ys \<restriction>\<^sub>l A)"
by (simp add: seq_filter_def)
lemma seq_filter_UNIV [simp]: "xs \<restriction>\<^sub>l UNIV = xs"
by (simp add: seq_filter_def)
lemma seq_filter_twice [simp]: "(xs \<restriction>\<^sub>l A) \<restriction>\<^sub>l B = xs \<restriction>\<^sub>l (A \<inter> B)"
by (simp add: seq_filter_def)
subsection \<open> Minus on lists \<close>
instantiation list :: (type) minus
begin
text \<open> We define list minus so that if the second list is not a prefix of the first, then an arbitrary
list longer than the combined length is produced. Thus we can always determined from the output
whether the minus is defined or not. \<close>
definition "xs - ys = (if (prefix ys xs) then drop (length ys) xs else [])"
instance ..
end
lemma minus_cancel [simp]: "xs - xs = []"
by (simp add: minus_list_def)
lemma append_minus [simp]: "(xs @ ys) - xs = ys"
by (simp add: minus_list_def)
lemma minus_right_nil [simp]: "xs - [] = xs"
by (simp add: minus_list_def)
lemma list_concat_minus_list_concat: "(s @ t) - (s @ z) = t - z"
by (simp add: minus_list_def)
lemma length_minus_list: "y \<le> x \<Longrightarrow> length(x - y) = length(x) - length(y)"
by (simp add: less_eq_list_def minus_list_def)
lemma map_list_minus:
"xs \<le> ys \<Longrightarrow> map f (ys - xs) = map f ys - map f xs"
by (simp add: drop_map less_eq_list_def map_mono_prefix minus_list_def)
lemma list_minus_first_tl [simp]:
"[x] \<le> xs \<Longrightarrow> (xs - [x]) = tl xs"
by (metis Prefix_Order.prefixE append.left_neutral append_minus list.sel(3) not_Cons_self2 tl_append2)
text \<open> Extra lemmas about @{term prefix} and @{term strict_prefix} \<close>
lemma prefix_concat_minus:
assumes "prefix xs ys"
shows "xs @ (ys - xs) = ys"
using assms by (metis minus_list_def prefix_drop)
lemma prefix_minus_concat:
assumes "prefix s t"
shows "(t - s) @ z = (t @ z) - s"
using assms by (simp add: Sublist.prefix_length_le minus_list_def)
lemma strict_prefix_minus_not_empty:
assumes "strict_prefix xs ys"
shows "ys - xs \<noteq> []"
using assms by (metis append_Nil2 prefix_concat_minus strict_prefix_def)
lemma strict_prefix_diff_minus:
assumes "prefix xs ys" and "xs \<noteq> ys"
shows "(ys - xs) \<noteq> []"
using assms by (simp add: strict_prefix_minus_not_empty)
lemma length_tl_list_minus_butlast_gt_zero:
assumes "length s < length t" and "strict_prefix (butlast s) t" and "length s > 0"
shows "length (tl (t - (butlast s))) > 0"
using assms
by (metis Nitpick.size_list_simp(2) butlast_snoc hd_Cons_tl length_butlast length_greater_0_conv length_tl less_trans nat_neq_iff strict_prefix_minus_not_empty prefix_order.dual_order.strict_implies_order prefix_concat_minus)
lemma list_minus_butlast_eq_butlast_list:
assumes "length t = length s" and "strict_prefix (butlast s) t"
shows "t - (butlast s) = [last t]"
using assms
by (metis append_butlast_last_id append_eq_append_conv butlast.simps(1) length_butlast less_numeral_extra(3) list.size(3) prefix_order.dual_order.strict_implies_order prefix_concat_minus prefix_length_less)
lemma butlast_strict_prefix_length_lt_imp_last_tl_minus_butlast_eq_last:
assumes "length s > 0" "strict_prefix (butlast s) t" "length s < length t"
shows "last (tl (t - (butlast s))) = (last t)"
using assms by (metis last_append last_tl length_tl_list_minus_butlast_gt_zero less_numeral_extra(3) list.size(3) append_minus strict_prefix_eq_exists)
lemma tl_list_minus_butlast_not_empty:
assumes "strict_prefix (butlast s) t" and "length s > 0" and "length t > length s"
shows "tl (t - (butlast s)) \<noteq> []"
using assms length_tl_list_minus_butlast_gt_zero by fastforce
lemma tl_list_minus_butlast_empty:
assumes "strict_prefix (butlast s) t" and "length s > 0" and "length t = length s"
shows "tl (t - (butlast s)) = []"
using assms by (simp add: list_minus_butlast_eq_butlast_list)
lemma concat_minus_list_concat_butlast_eq_list_minus_butlast:
assumes "prefix (butlast u) s"
shows "(t @ s) - (t @ (butlast u)) = s - (butlast u)"
using assms by (metis append_assoc prefix_concat_minus append_minus)
lemma tl_list_minus_butlast_eq_empty:
assumes "strict_prefix (butlast s) t" and "length s = length t"
shows "tl (t - (butlast s)) = []"
using assms by (metis list.sel(3) list_minus_butlast_eq_butlast_list)
(* this can be shown using length_tl, but care is needed when list is empty? *)
lemma prefix_length_tl_minus:
assumes "strict_prefix s t"
shows "length (tl (t-s)) = (length (t-s)) - 1"
by (auto)
lemma length_list_minus:
assumes "strict_prefix s t"
shows "length(t - s) = length(t) - length(s)"
using assms by (simp add: minus_list_def prefix_order.dual_order.strict_implies_order)
lemma length_minus_le: "length (ys - xs) \<le> length ys"
by (simp add: minus_list_def)
lemma length_minus_less: "\<lbrakk> xs \<le> ys; xs \<noteq> [] \<rbrakk> \<Longrightarrow> length (ys - xs) < length ys"
by (auto simp add: minus_list_def less_eq_list_def)
(metis diff_less length_greater_0_conv prefix_bot.extremum_uniqueI)
lemma filter_minus [simp]: "ys \<le> xs \<Longrightarrow> filter P (xs - ys) = filter P xs - filter P ys"
by (simp add: minus_list_def less_eq_list_def filter_mono_prefix)
(metis filter_append filter_mono_prefix prefix_drop same_append_eq)
subsection \<open> Laws on @{term list_update} \<close>
lemma list_update_0: "length(xs) > 0 \<Longrightarrow> xs[0 := x] = x # tl xs"
by (metis length_0_conv list.collapse list_update_code(2) nat_less_le)
lemma tl_list_update: "\<lbrakk> length xs > 0; k > 0 \<rbrakk> \<Longrightarrow> tl(xs[k := v]) = (tl xs)[k-1 := v]"
by (metis One_nat_def Suc_pred length_greater_0_conv list.collapse list.sel(3) list_update_code(3))
subsection \<open> Laws on @{term take}, @{term drop}, and @{term nths} \<close>
lemma take_prefix: "m \<le> n \<Longrightarrow> take m xs \<le> take n xs"
by (metis Prefix_Order.prefixI append_take_drop_id min_absorb2 take_append take_take)
lemma nths_atLeastAtMost_0_take: "nths xs {0..m} = take (Suc m) xs"
by (metis atLeast0AtMost lessThan_Suc_atMost nths_upt_eq_take)
lemma nths_atLeastLessThan_0_take: "nths xs {0..<m} = take m xs"
by (simp add: atLeast0LessThan)
lemma nths_atLeastAtMost_prefix: "m \<le> n \<Longrightarrow> nths xs {0..m} \<le> nths xs {0..n}"
by (simp add: nths_atLeastAtMost_0_take take_prefix)
lemma sorted_nths_atLeastAtMost_0: "\<lbrakk> m \<le> n; sorted (nths xs {0..n}) \<rbrakk> \<Longrightarrow> sorted (nths xs {0..m})"
using nths_atLeastAtMost_prefix sorted_prefix by blast
lemma sorted_nths_atLeastLessThan_0: "\<lbrakk> m \<le> n; sorted (nths xs {0..<n}) \<rbrakk> \<Longrightarrow> sorted (nths xs {0..<m})"
by (metis atLeast0LessThan nths_upt_eq_take sorted_prefix take_prefix)
lemma list_augment_as_update:
"k < length xs \<Longrightarrow> list_augment xs k x = list_update xs k x"
by (metis list_augment_def list_augment_idem list_update_overwrite)
lemma nths_list_update_out: "k \<notin> A \<Longrightarrow> nths (list_update xs k x) A = nths xs A"
apply (induct xs arbitrary: k x A)
apply (auto)
apply (rename_tac a xs k x A)
apply (case_tac k)
apply (auto simp add: nths_Cons)
done
lemma nths_list_augment_out: "\<lbrakk> k < length xs; k \<notin> A \<rbrakk> \<Longrightarrow> nths (list_augment xs k x) A = nths xs A"
by (simp add: list_augment_as_update nths_list_update_out)
lemma nths_none: "\<forall>i \<in> I. i \<ge> length xs \<Longrightarrow> nths xs I = []"
apply (simp add: nths_def)
apply (subst filter_False)
apply (metis atLeastLessThan_iff in_set_zip leD nth_mem set_upt)
apply simp
done
lemma nths_uptoLessThan:
"\<lbrakk> m \<le> n; n < length xs \<rbrakk> \<Longrightarrow> nths xs {m..n} = xs ! m # nths xs {Suc m..n}"
proof (induct xs arbitrary: m n)
case Nil
then show ?case by (simp)
next
case (Cons a xs)
have l1: "\<And> m n. \<lbrakk> 0 < m; m \<le> n \<rbrakk> \<Longrightarrow> {j. m \<le> Suc j \<and> Suc j \<le> n} = {m-1..n-1}"
by (auto)
have l2: "\<And> m n. \<lbrakk> 0 < m; m \<le> n \<rbrakk> \<Longrightarrow> {j. m \<le> j \<and> Suc j \<le> n} = {m..n-1}"
by (auto)
from Cons show ?case by (auto simp add: nths_Cons l1 l2)
qed
lemma nths_upt_nth: "\<lbrakk> j < i; i < length xs \<rbrakk> \<Longrightarrow> (nths xs {0..<i}) ! j = xs ! j"
by (metis lessThan_atLeast0 nth_take nths_upt_eq_take)
lemma nths_upt_length: "\<lbrakk> m \<le> n; n \<le> length xs \<rbrakk> \<Longrightarrow> length (nths xs {m..<n}) = n-m"
by (metis atLeastLessThan_empty diff_is_0_eq length_map length_upt list.size(3) not_less nths_empty seq_extract_as_map seq_extract_def)
lemma nths_upt_le_length:
"\<lbrakk> m \<le> n; Suc n \<le> length xs \<rbrakk> \<Longrightarrow> length (nths xs {m..n}) = Suc n - m"
by (metis atLeastLessThanSuc_atLeastAtMost le_SucI nths_upt_length)
lemma sl1: "n > 0 \<Longrightarrow> {j. Suc j \<le> n} = {0..n-1}"
by (auto)
lemma sl2: "\<lbrakk> 0 < m; m \<le> n \<rbrakk> \<Longrightarrow> {j. m \<le> Suc j \<and> Suc j \<le> n} = {m-1..n-1}"
by auto
lemma nths_upt_le_nth: "\<lbrakk> m \<le> n; Suc n \<le> length xs; i < Suc n - m \<rbrakk>
\<Longrightarrow> (nths xs {m..n}) ! i = xs ! (i + m)"