-
Notifications
You must be signed in to change notification settings - Fork 0
/
BPlusTree_Set.thy
3666 lines (3393 loc) · 150 KB
/
BPlusTree_Set.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
theory BPlusTree_Set
imports
BPlusTree_Split
"HOL-Data_Structures.Set_Specs"
begin
section "Set interpretation"
(* obsolete fact *)
lemma insert_list_length[simp]:
assumes "sorted_less ks"
and "set (insert_list k ks) = set ks \<union> {k}"
and "sorted_less ks \<Longrightarrow> sorted_less (insert_list k ks)"
shows "length (insert_list k ks) = length ks + (if k \<in> set ks then 0 else 1)"
proof -
have "distinct (insert_list k ks)"
using assms(1) assms(3) strict_sorted_iff by blast
then have "length (insert_list k ks) = card (set (insert_list k ks))"
by (simp add: distinct_card)
also have "\<dots> = card (set ks \<union> {k})"
using assms(2) by presburger
also have "\<dots> = card (set ks) + (if k \<in> set ks then 0 else 1)"
by (cases "k \<in> set ks") (auto simp add: insert_absorb)
also have "\<dots> = length ks + (if k \<in> set ks then 0 else 1)"
using assms(1) distinct_card strict_sorted_iff by auto
finally show ?thesis.
qed
lemma delete_list_length[simp]:
assumes "sorted_less ks"
and "set (delete_list k ks) = set ks - {k}"
and "sorted_less ks \<Longrightarrow> sorted_less (delete_list k ks)"
shows "length (delete_list k ks) = length ks - (if k \<in> set ks then 1 else 0)"
proof -
have "distinct (delete_list k ks)"
using assms(1) assms(3) strict_sorted_iff by blast
then have "length (delete_list k ks) = card (set (delete_list k ks))"
by (simp add: distinct_card)
also have "\<dots> = card (set ks - {k})"
using assms(2) by presburger
also have "\<dots> = card (set ks) - (if k \<in> set ks then 1 else 0)"
by (cases "k \<in> set ks") (auto)
also have "\<dots> = length ks - (if k \<in> set ks then 1 else 0)"
by (metis assms(1) distinct_card strict_sorted_iff)
finally show ?thesis.
qed
lemma ins_list_length[simp]:
assumes "sorted_less ks"
shows "length (ins_list k ks) = length ks + (if k \<in> set ks then 0 else 1)"
using insert_list_length[of ks ins_list k]
by (simp add: assms set_ins_list sorted_ins_list)
lemma del_list_length[simp]:
assumes "sorted_less ks"
shows "length (del_list k ks) = length ks - (if k \<in> set ks then 1 else 0)"
using delete_list_length[of ks ins_list k]
by (simp add: assms set_del_list sorted_del_list)
(* TODO what if we define a function "list_split" that returns
a split list for mapping arbitrary f (separators) and g (subtrees)
s.th. f :: 'a \<Rightarrow> ('b::linorder) and g :: 'a \<Rightarrow> 'a bplustree
this would allow for key,pointer pairs to be inserted into the tree *)
(* TODO what if the keys are the pointers? *)
locale split_set = split_tree: split_tree split
for split::
"('a bplustree \<times> 'a::{linorder,order_top}) list \<Rightarrow> 'a
\<Rightarrow> ('a bplustree \<times> 'a) list \<times> ('a bplustree \<times> 'a) list" +
fixes isin_list :: "'a \<Rightarrow> ('a::{linorder,order_top}) list \<Rightarrow> bool"
and insert_list :: "'a \<Rightarrow> ('a::{linorder,order_top}) list \<Rightarrow> 'a list"
and delete_list :: "'a \<Rightarrow> ('a::{linorder,order_top}) list \<Rightarrow> 'a list"
assumes insert_list_req:
(* TODO locale that derives such a function from a split function similar to the above *)
"sorted_less ks \<Longrightarrow> isin_list x ks = (x \<in> set ks)"
"sorted_less ks \<Longrightarrow> insert_list x ks = ins_list x ks"
"sorted_less ks \<Longrightarrow> delete_list x ks = del_list x ks"
begin
lemmas split_req = split_tree.split_req
lemmas split_conc = split_tree.split_req(1)
lemmas split_sorted = split_tree.split_req(2,3)
lemma insert_list_length[simp]:
assumes "sorted_less ks"
shows "length (insert_list k ks) = length ks + (if k \<in> set ks then 0 else 1)"
using insert_list_req
by (simp add: assms)
lemma set_insert_list[simp]:
"sorted_less ks \<Longrightarrow> set (insert_list k ks) = set ks \<union> {k}"
by (simp add: insert_list_req set_ins_list)
lemma sorted_insert_list[simp]:
"sorted_less ks \<Longrightarrow> sorted_less (insert_list k ks)"
by (simp add: insert_list_req sorted_ins_list)
lemma delete_list_length[simp]:
assumes "sorted_less ks"
shows "length (delete_list k ks) = length ks - (if k \<in> set ks then 1 else 0)"
using insert_list_req
by (simp add: assms)
lemma set_delete_list[simp]:
"sorted_less ks \<Longrightarrow> set (delete_list k ks) = set ks - {k}"
by (simp add: insert_list_req set_del_list)
lemma sorted_delete_list[simp]:
"sorted_less ks \<Longrightarrow> sorted_less (delete_list k ks)"
by (simp add: insert_list_req sorted_del_list)
definition "empty_bplustree = (Leaf [])"
subsection "Membership"
fun isin:: "'a bplustree \<Rightarrow> 'a \<Rightarrow> bool" where
"isin (Leaf ks) x = (isin_list x ks)" |
"isin (Node ts t) x = (
case split ts x of (_,(sub,sep)#rs) \<Rightarrow> (
isin sub x
)
| (_,[]) \<Rightarrow> isin t x
)"
text "Isin proof"
thm isin_simps
(* copied from comment in List_Ins_Del *)
lemma sorted_ConsD: "sorted_less (y # xs) \<Longrightarrow> x \<le> y \<Longrightarrow> x \<notin> set xs"
by (auto simp: sorted_Cons_iff)
lemma sorted_snocD: "sorted_less (xs @ [y]) \<Longrightarrow> y \<le> x \<Longrightarrow> x \<notin> set xs"
by (auto simp: sorted_snoc_iff)
lemmas isin_simps2 = sorted_lems sorted_ConsD sorted_snocD
(*-----------------------------*)
lemma isin_sorted: "sorted_less (xs@a#ys) \<Longrightarrow>
(x \<in> set (xs@a#ys)) = (if x < a then x \<in> set xs else x \<in> set (a#ys))"
by (auto simp: isin_simps2)
(* lift to split *)
lemma isin_sorted_split:
assumes "Laligned (Node ts t) u"
and "sorted_less (leaves (Node ts t))"
and "split ts x = (ls, rs)"
shows "x \<in> set (leaves (Node ts t)) = (x \<in> set (leaves_list rs @ leaves t))"
proof (cases ls)
case Nil
then have "ts = rs"
using assms by (auto dest!: split_conc)
then show ?thesis by simp
next
case Cons
then obtain ls' sub sep where ls_tail_split: "ls = ls' @ [(sub,sep)]"
by (metis list.simps(3) rev_exhaust surj_pair)
then have x_sm_sep: "sep < x"
using split_req(2)[of ts x ls' sub sep rs]
using Laligned_sorted_separators[OF assms(1)]
using assms sorted_cons sorted_snoc
by blast
moreover have leaves_split: "leaves (Node ts t) = leaves_list ls @ leaves_list rs @ leaves t"
using assms(3) split_tree.leaves_split by blast
then show ?thesis
proof (cases "leaves_list ls")
case Nil
then show ?thesis
using leaves_split by auto
next
case Cons
then obtain leavesls' l' where leaves_tail_split: "leaves_list ls = leavesls' @ [l']"
by (metis list.simps(3) rev_exhaust)
then have "l' \<le> sep"
proof -
have "l' \<in> set (leaves_list ls)"
using leaves_tail_split by force
then have "l' \<in> set (leaves (Node ls' sub))"
using ls_tail_split
by auto
moreover have "Laligned (Node ls' sub) sep"
using assms split_conc[OF assms(3)] Cons ls_tail_split
using Laligned_split_left[of ls' sub sep rs t u]
by simp
ultimately show ?thesis
using Laligned_leaves_inbetween[of "Node ls' sub" sep]
by blast
qed
then show ?thesis
using assms(2) ls_tail_split leaves_tail_split leaves_split x_sm_sep
using isin_sorted[of "leavesls'" l' "leaves_list rs @ leaves t" x]
by auto
qed
qed
lemma isin_sorted_split_right:
assumes "split ts x = (ls, (sub,sep)#rs)"
and "sorted_less (leaves (Node ts t))"
and "Laligned (Node ts t) u"
shows "x \<in> set (leaves_list ((sub,sep)#rs) @ leaves t) = (x \<in> set (leaves sub))"
proof -
from assms have "x \<le> sep"
proof -
from assms have "sorted_less (separators ts)"
by (meson Laligned_sorted_inorder sorted_cons sorted_inorder_separators sorted_snoc)
then show ?thesis
using split_req(3)
using assms
by fastforce
qed
moreover have leaves_split: "leaves (Node ts t) = leaves_list ls @ leaves sub @ leaves_list rs @ leaves t"
using split_conc[OF assms(1)] by auto
ultimately show ?thesis
proof (cases "leaves_list rs @ leaves t")
case Nil
then show ?thesis
using leaves_split by auto
next
case (Cons r' rs')
then have "sep < r'"
by (metis Laligned_split_right aligned_leaves_inbetween assms(1) assms(3) leaves.simps(2) list.set_intros(1) split_set.split_conc split_set_axioms)
then have "x < r'"
using \<open>x \<le> sep\<close> by auto
moreover have "sorted_less (leaves_list ((sub,sep)#rs) @ leaves t)"
using assms sorted_wrt_append split_conc
by fastforce
ultimately show ?thesis
using isin_sorted[of "leaves sub" "r'" "rs'" x] Cons
by auto
qed
qed
theorem isin_set_inorder:
assumes "sorted_less (leaves t)"
and "aligned l t u"
shows "isin t x = (x \<in> set (leaves t))"
using assms
proof(induction t x arbitrary: l u rule: isin.induct)
case (2 ts t x)
then obtain ls rs where list_split: "split ts x = (ls, rs)"
by (meson surj_pair)
then have list_conc: "ts = ls @ rs"
using split_conc by auto
show ?case
proof (cases rs)
case Nil
then have "isin (Node ts t) x = isin t x"
by (simp add: list_split)
also have "\<dots> = (x \<in> set (leaves t))"
using "2.IH"(1)[of ls rs] list_split Nil
using "2.prems" sorted_leaves_induct_last align_last'
by metis
also have "\<dots> = (x \<in> set (leaves (Node ts t)))"
using isin_sorted_split
using "2.prems" list_split list_conc Nil
by (metis aligned_imp_Laligned leaves.simps(2) leaves_conc same_append_eq self_append_conv)
finally show ?thesis .
next
case (Cons a list)
then obtain sub sep where a_split: "a = (sub,sep)"
by (cases a)
then have "isin (Node ts t) x = isin sub x"
using list_split Cons a_split
by auto
also have "\<dots> = (x \<in> set (leaves sub))"
using "2.IH"(2)[of ls rs "(sub,sep)" list sub sep]
using "2.prems" a_split list_conc list_split local.Cons sorted_leaves_induct_subtree
align_sub
by (metis in_set_conv_decomp)
also have "\<dots> = (x \<in> set (leaves (Node ts t)))"
using isin_sorted_split
using isin_sorted_split_right "2.prems" list_split Cons a_split
using aligned_imp_Laligned by blast
finally show ?thesis .
qed
qed (auto simp add: insert_list_req)
theorem isin_set_Linorder:
assumes "sorted_less (leaves t)"
and "Laligned t u"
shows "isin t x = (x \<in> set (leaves t))"
using assms
proof(induction t x arbitrary: u rule: isin.induct)
case (2 ts t x)
then obtain ls rs where list_split: "split ts x = (ls, rs)"
by (meson surj_pair)
then have list_conc: "ts = ls @ rs"
using split_conc by auto
show ?case
proof (cases rs)
case Nil
then have "isin (Node ts t) x = isin t x"
by (simp add: list_split)
also have "\<dots> = (x \<in> set (leaves t))"
by (metis "2.IH"(1) "2.prems"(1) "2.prems"(2) Lalign_Llast list_split local.Nil sorted_leaves_induct_last)
also have "\<dots> = (x \<in> set (leaves (Node ts t)))"
using isin_sorted_split
using "2.prems" list_split list_conc Nil
by simp
finally show ?thesis .
next
case (Cons a list)
then obtain sub sep where a_split: "a = (sub,sep)"
by (cases a)
then have "isin (Node ts t) x = isin sub x"
using list_split Cons a_split
by auto
also have "\<dots> = (x \<in> set (leaves sub))"
using "2.IH"(2)[of ls rs "(sub,sep)" list sub sep]
using "2.prems" a_split list_conc list_split local.Cons sorted_leaves_induct_subtree
align_sub
by (metis Lalign_Llast Laligned_split_left)
also have "\<dots> = (x \<in> set (leaves (Node ts t)))"
using isin_sorted_split
using isin_sorted_split_right "2.prems" list_split Cons a_split
by simp
finally show ?thesis .
qed
qed (auto simp add: insert_list_req)
corollary isin_set_Linorder_top:
assumes "sorted_less (leaves t)"
and "Laligned t top"
shows "isin t x = (x \<in> set (leaves t))"
using assms isin_set_Linorder
by simp
subsection "Insertion"
text "The insert function requires an auxiliary data structure
and auxiliary invariant functions."
datatype 'b up\<^sub>i = T\<^sub>i "'b bplustree" | Up\<^sub>i "'b bplustree" 'b "'b bplustree"
fun order_up\<^sub>i where
"order_up\<^sub>i k (T\<^sub>i sub) = order k sub" |
"order_up\<^sub>i k (Up\<^sub>i l a r) = (order k l \<and> order k r)"
fun root_order_up\<^sub>i where
"root_order_up\<^sub>i k (T\<^sub>i sub) = root_order k sub" |
"root_order_up\<^sub>i k (Up\<^sub>i l a r) = (order k l \<and> order k r)"
fun height_up\<^sub>i where
"height_up\<^sub>i (T\<^sub>i t) = height t" |
"height_up\<^sub>i (Up\<^sub>i l a r) = max (height l) (height r)"
fun bal_up\<^sub>i where
"bal_up\<^sub>i (T\<^sub>i t) = bal t" |
"bal_up\<^sub>i (Up\<^sub>i l a r) = (height l = height r \<and> bal l \<and> bal r)"
fun inorder_up\<^sub>i where
"inorder_up\<^sub>i (T\<^sub>i t) = inorder t" |
"inorder_up\<^sub>i (Up\<^sub>i l a r) = inorder l @ [a] @ inorder r"
fun leaves_up\<^sub>i where
"leaves_up\<^sub>i (T\<^sub>i t) = leaves t" |
"leaves_up\<^sub>i (Up\<^sub>i l a r) = leaves l @ leaves r"
fun aligned_up\<^sub>i where
"aligned_up\<^sub>i l (T\<^sub>i t) u = aligned l t u" |
"aligned_up\<^sub>i l (Up\<^sub>i lt a rt) u = (aligned l lt a \<and> aligned a rt u)"
fun Laligned_up\<^sub>i where
"Laligned_up\<^sub>i (T\<^sub>i t) u = Laligned t u" |
"Laligned_up\<^sub>i (Up\<^sub>i lt a rt) u = (Laligned lt a \<and> aligned a rt u)"
text "The following function merges two nodes and returns separately split nodes
if an overflow occurs"
(* note here that splitting away the last element is actually very nice
from the implementation perspective *)
fun node\<^sub>i:: "nat \<Rightarrow> ('a bplustree \<times> 'a) list \<Rightarrow> 'a bplustree \<Rightarrow> 'a up\<^sub>i" where
"node\<^sub>i k ts t = (
if length ts \<le> 2*k then T\<^sub>i (Node ts t)
else (
case split_half ts of (ls, rs) \<Rightarrow>
case last ls of (sub,sep) \<Rightarrow>
Up\<^sub>i (Node (butlast ls) sub) sep (Node rs t)
)
)"
fun Lnode\<^sub>i:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a up\<^sub>i" where
"Lnode\<^sub>i k ts = (
if length ts \<le> 2*k then T\<^sub>i (Leaf ts)
else (
case split_half ts of (ls, rs) \<Rightarrow>
Up\<^sub>i (Leaf ls) (last ls) (Leaf rs)
)
)"
fun ins:: "nat \<Rightarrow> 'a \<Rightarrow> 'a bplustree \<Rightarrow> 'a up\<^sub>i" where
"ins k x (Leaf ks) = Lnode\<^sub>i k (insert_list x ks)" |
"ins k x (Node ts t) = (
case split ts x of
(ls,(sub,sep)#rs) \<Rightarrow>
(case ins k x sub of
Up\<^sub>i l a r \<Rightarrow>
node\<^sub>i k (ls@(l,a)#(r,sep)#rs) t |
T\<^sub>i a \<Rightarrow>
T\<^sub>i (Node (ls@(a,sep)#rs) t)) |
(ls, []) \<Rightarrow>
(case ins k x t of
Up\<^sub>i l a r \<Rightarrow>
node\<^sub>i k (ls@[(l,a)]) r |
T\<^sub>i a \<Rightarrow>
T\<^sub>i (Node ls a)
)
)"
fun tree\<^sub>i::"'a up\<^sub>i \<Rightarrow> 'a bplustree" where
"tree\<^sub>i (T\<^sub>i sub) = sub" |
"tree\<^sub>i (Up\<^sub>i l a r) = (Node [(l,a)] r)"
fun insert::"nat \<Rightarrow> 'a \<Rightarrow> 'a bplustree \<Rightarrow> 'a bplustree" where
"insert k x t = tree\<^sub>i (ins k x t)"
subsection "Proofs of functional correctness"
lemma nodei_ti_simp: "node\<^sub>i k ts t = T\<^sub>i x \<Longrightarrow> x = Node ts t"
apply (cases "length ts \<le> 2*k")
apply (auto split!: list.splits prod.splits)
done
lemma Lnodei_ti_simp: "Lnode\<^sub>i k ts = T\<^sub>i x \<Longrightarrow> x = Leaf ts"
apply (cases "length ts \<le> 2*k")
apply (auto split!: list.splits)
done
lemma split_set:
assumes "split ts z = (ls,(a,b)#rs)"
shows "(a,b) \<in> set ts"
and "(x,y) \<in> set ls \<Longrightarrow> (x,y) \<in> set ts"
and "(x,y) \<in> set rs \<Longrightarrow> (x,y) \<in> set ts"
and "set ls \<union> set rs \<union> {(a,b)} = set ts"
and "\<exists>x \<in> set ts. b \<in> Basic_BNFs.snds x"
using split_conc assms by fastforce+
lemma split_length:
"split ts x = (ls, rs) \<Longrightarrow> length ls + length rs = length ts"
by (auto dest: split_conc)
(* TODO way to use this for custom case distinction? *)
lemma node\<^sub>i_cases: "length xs \<le> k \<or> (\<exists>ls sub sep rs. split_half xs = (ls@[(sub,sep)],rs))"
proof -
have "\<not> length xs \<le> k \<Longrightarrow> length xs \<ge> 1"
by linarith
then show ?thesis
using split_half_not_empty
by fastforce
qed
lemma Lnode\<^sub>i_cases: "length xs \<le> k \<or> (\<exists>ls sep rs. split_half xs = (ls@[sep],rs))"
proof -
have "\<not> length xs \<le> k \<Longrightarrow> length xs \<ge> 1"
by linarith
then show ?thesis
using split_half_not_empty
by fastforce
qed
lemma root_order_tree\<^sub>i: "root_order_up\<^sub>i (Suc k) t = root_order (Suc k) (tree\<^sub>i t)"
apply (cases t)
apply auto
done
lemma length_take_left: "length (take ((length ts + 1) div 2) ts) = (length ts + 1) div 2"
apply (cases ts)
apply auto
done
lemma node\<^sub>i_root_order:
assumes "length ts > 0"
and "length ts \<le> 4*k+1"
and "\<forall>x \<in> set (subtrees ts). order k x"
and "order k t"
shows "root_order_up\<^sub>i k (node\<^sub>i k ts t)"
proof (cases "length ts \<le> 2*k")
case True
then show ?thesis
using assms
by (simp add: node\<^sub>i.simps)
next
case False
then obtain ls sub sep rs where split_half_ts:
"take ((length ts + 1) div 2) ts = ls@[(sub,sep)]"
using split_half_not_empty[of ts]
by auto
then have length_ls: "length ls = (length ts + 1) div 2 - 1"
by (metis One_nat_def add_diff_cancel_right' add_self_div_2 bits_1_div_2 length_append length_take_left list.size(3) list.size(4) odd_one odd_succ_div_two)
also have "\<dots> \<le> (4*k + 1) div 2"
using assms(2) by simp
also have "\<dots> = 2*k"
by auto
finally have "length ls \<le> 2*k"
by simp
moreover have "length ls \<ge> k"
using False length_ls by simp
moreover have "set (ls@[(sub,sep)]) \<subseteq> set ts"
by (metis split_half_ts(1) set_take_subset)
ultimately have o_r: "order k (Node ls sub)"
using split_half_ts assms by auto
have
"butlast (take ((length ts + 1) div 2) ts) = ls"
"last (take ((length ts + 1) div 2) ts) = (sub,sep)"
using split_half_ts by auto
then show ?thesis
using o_r assms set_drop_subset[of _ ts]
by (auto simp add: False split_half_ts split!: prod.splits)
qed
lemma node\<^sub>i_order_helper:
assumes "length ts \<ge> k"
and "length ts \<le> 4*k+1"
and "\<forall>x \<in> set (subtrees ts). order k x"
and "order k t"
shows "case (node\<^sub>i k ts t) of T\<^sub>i t \<Rightarrow> order k t | _ \<Rightarrow> True"
proof (cases "length ts \<le> 2*k")
case True
then show ?thesis
using assms
by (simp add: node\<^sub>i.simps)
next
case False
then obtain sub sep ls where
"take ((length ts + 1) div 2) ts = ls@[(sub,sep)]"
using split_half_not_empty[of ts]
by fastforce
then show ?thesis
using assms by simp
qed
lemma node\<^sub>i_order:
assumes "length ts \<ge> k"
and "length ts \<le> 4*k+1"
and "\<forall>x \<in> set (subtrees ts). order k x"
and "order k t"
shows "order_up\<^sub>i k (node\<^sub>i k ts t)"
apply(cases "node\<^sub>i k ts t")
using node\<^sub>i_root_order node\<^sub>i_order_helper assms apply fastforce
by (metis (full_types) assms le_0_eq nat_le_linear node\<^sub>i.elims node\<^sub>i_root_order order_up\<^sub>i.simps(2) root_order_up\<^sub>i.simps(2) up\<^sub>i.simps(4) verit_comp_simplify1(3))
lemma Lnode\<^sub>i_root_order:
assumes "length ts > 0"
and "length ts \<le> 4*k"
shows "root_order_up\<^sub>i k (Lnode\<^sub>i k ts)"
proof (cases "length ts \<le> 2*k")
case True
then show ?thesis
using assms
by simp
next
case False
then obtain ls sep rs where split_half_ts:
"take ((length ts + 1) div 2) ts = ls@[sep]"
"drop ((length ts + 1) div 2) ts = rs"
using split_half_not_empty[of ts]
by auto
then have length_ls: "length ls = ((length ts + 1) div 2) - 1"
by (metis One_nat_def add_diff_cancel_right' add_self_div_2 bits_1_div_2 length_append length_take_left list.size(3) list.size(4) odd_one odd_succ_div_two)
also have "\<dots> < (4*k + 1) div 2"
using assms(2)
by (smt (z3) Groups.add_ac(2) One_nat_def split_half_ts add.right_neutral diff_is_0_eq' div_le_mono le_add_diff_inverse le_neq_implies_less length_append length_take_left less_add_Suc1 less_imp_diff_less list.size(4) nat_le_linear not_less_eq plus_nat.simps(2))
also have "\<dots> = 2*k"
by auto
finally have "length ls < 2*k"
by simp
moreover have "length ls \<ge> k"
using False length_ls by simp
ultimately have o_l: "order k (Leaf (ls@[sep]))"
using set_take_subset assms split_half_ts
by fastforce
then show ?thesis
using assms split_half_ts False
by auto
qed
lemma Lnode\<^sub>i_order_helper:
assumes "length ts \<ge> k"
and "length ts \<le> 4*k+1"
shows "case (Lnode\<^sub>i k ts) of T\<^sub>i t \<Rightarrow> order k t | _ \<Rightarrow> True"
proof (cases "length ts \<le> 2*k")
case True
then show ?thesis
using assms
by (simp add: node\<^sub>i.simps)
next
case False
then obtain sep ls where
"take ((length ts + 1) div 2) ts = ls@[sep]"
using split_half_not_empty[of ts]
by fastforce
then show ?thesis
using assms by simp
qed
lemma Lnode\<^sub>i_order:
assumes "length ts \<ge> k"
and "length ts \<le> 4*k"
shows "order_up\<^sub>i k (Lnode\<^sub>i k ts)"
apply(cases "Lnode\<^sub>i k ts")
apply (metis Lnode\<^sub>i_order_helper One_nat_def add.right_neutral add_Suc_right assms(1) assms(2) le_imp_less_Suc less_le order_up\<^sub>i.simps(1) up\<^sub>i.simps(5))
by (metis Lnode\<^sub>i.elims Lnode\<^sub>i_root_order assms(1) assms(2) diff_is_0_eq' le_0_eq le_add_diff_inverse mult_2 order_up\<^sub>i.simps(2) root_order_up\<^sub>i.simps(2) up\<^sub>i.simps(3) verit_comp_simplify1(3))
(* explicit proof *)
lemma ins_order:
"k > 0 \<Longrightarrow> sorted_less (leaves t) \<Longrightarrow> order k t \<Longrightarrow> order_up\<^sub>i k (ins k x t)"
proof(induction k x t rule: ins.induct)
case (1 k x ts)
then show ?case
by auto (* this proof requires both sorted_less and k > 0 *)
next
case (2 k x ts t)
then obtain ls rs where split_res: "split ts x = (ls, rs)"
by (meson surj_pair)
then have split_app: "ts = ls@rs"
using split_conc
by simp
show ?case
proof (cases rs)
case Nil
then have "order_up\<^sub>i k (ins k x t)"
using 2 split_res sorted_leaves_induct_last
by auto
then show ?thesis
using Nil 2 split_app split_res Nil node\<^sub>i_order
by (auto split!: up\<^sub>i.splits simp del: node\<^sub>i.simps)
next
case (Cons a list)
then obtain sub sep where a_prod: "a = (sub, sep)"
by (cases a)
then have "sorted_less (leaves sub)"
using "2"(4) Cons sorted_leaves_induct_subtree split_app
by blast
then have "order_up\<^sub>i k (ins k x sub)"
using "2.IH"(2) "2.prems" a_prod local.Cons split_app split_res
by auto
then show ?thesis
using 2 split_app Cons length_append node\<^sub>i_order[of k "ls@_#_#list"] a_prod split_res
by (auto split!: up\<^sub>i.splits simp del: node\<^sub>i.simps simp add: order_impl_root_order)
qed
qed
(* notice this is almost a duplicate of ins_order *)
lemma ins_root_order:
assumes "k > 0" "sorted_less (leaves t)" "root_order k t"
shows "root_order_up\<^sub>i k (ins k x t)"
proof(cases t)
case (Leaf ks)
then show ?thesis
using assms by (auto simp add: Lnode\<^sub>i_order min_absorb2) (* this proof requires both sorted_less and k > 0 *)
next
case (Node ts t)
then obtain ls rs where split_res: "split ts x = (ls, rs)"
by (meson surj_pair)
then have split_app: "ls@rs = ts"
using split_conc
by fastforce
show ?thesis
proof (cases rs)
case Nil
then have "order_up\<^sub>i k (ins k x t)"
using Node assms split_res sorted_leaves_induct_last
using ins_order[of k t]
by auto
then show ?thesis
using Nil Node split_app split_res assms node\<^sub>i_root_order
by (auto split!: up\<^sub>i.splits simp del: node\<^sub>i.simps simp add: order_impl_root_order)
next
case (Cons a list)
then obtain sub sep where a_prod: "a = (sub, sep)"
by (cases a)
then have "sorted_less (leaves sub)"
using Node assms(2) local.Cons sorted_leaves_induct_subtree split_app
by blast
then have "order_up\<^sub>i k (ins k x sub)"
using Node a_prod assms ins_order local.Cons split_app
by auto
then show ?thesis
using assms split_app Cons length_append Node node\<^sub>i_root_order a_prod split_res
by (auto split!: up\<^sub>i.splits simp del: node\<^sub>i.simps simp add: order_impl_root_order)
qed
qed
lemma height_list_split: "height_up\<^sub>i (Up\<^sub>i (Node ls a) b (Node rs t)) = height (Node (ls@(a,b)#rs) t) "
by (induction ls) (auto simp add: max.commute)
lemma node\<^sub>i_height: "height_up\<^sub>i (node\<^sub>i k ts t) = height (Node ts t)"
proof(cases "length ts \<le> 2*k")
case False
then obtain ls sub sep rs where
split_half_ts: "split_half ts = (ls@[(sub,sep)], rs)"
by (meson node\<^sub>i_cases)
then have "node\<^sub>i k ts t = Up\<^sub>i (Node ls (sub)) sep (Node rs t)"
using False by simp
then have "height_up\<^sub>i (node\<^sub>i k ts t) = height (Node (ls@(sub,sep)#rs) t)"
by (metis height_list_split)
also have "\<dots> = height (Node ts t)"
by (metis (no_types, lifting) Pair_inject append_Cons append_eq_append_conv2 append_take_drop_id self_append_conv split_half.simps split_half_ts)
finally show ?thesis.
qed simp
lemma Lnode\<^sub>i_height: "height_up\<^sub>i (Lnode\<^sub>i k xs) = height (Leaf xs)"
by (auto)
lemma bal_up\<^sub>i_tree: "bal_up\<^sub>i t = bal (tree\<^sub>i t)"
apply(cases t)
apply auto
done
lemma bal_list_split: "bal (Node (ls@(a,b)#rs) t) \<Longrightarrow> bal_up\<^sub>i (Up\<^sub>i (Node ls a) b (Node rs t))"
by (auto simp add: image_constant_conv)
lemma node\<^sub>i_bal:
assumes "bal (Node ts t)"
shows "bal_up\<^sub>i (node\<^sub>i k ts t)"
using assms
proof(cases "length ts \<le> 2*k")
case False
then obtain ls sub sep rs where
split_half_ts: "split_half ts = (ls@[(sub,sep)], rs)"
by (meson node\<^sub>i_cases)
then have "bal (Node (ls@(sub,sep)#rs) t)"
using assms append_take_drop_id[where n="(length ts + 1) div 2" and xs=ts]
by auto
then show ?thesis
using split_half_ts assms False
by (auto simp del: bal.simps bal_up\<^sub>i.simps dest!: bal_list_split[of ls sub sep rs t])
qed simp
lemma node\<^sub>i_aligned:
assumes "aligned l (Node ts t) u"
shows "aligned_up\<^sub>i l (node\<^sub>i k ts t) u"
using assms
proof (cases "length ts \<le> 2*k")
case False
then obtain ls sub sep rs where
split_half_ts: "split_half ts = (ls@[(sub,sep)], rs)"
by (meson node\<^sub>i_cases)
then have "aligned l (Node ls sub) sep"
by (metis aligned_split_left append.assoc append_Cons append_take_drop_id assms prod.sel(1) split_half.simps)
moreover have "aligned sep (Node rs t) u"
by (smt (z3) Pair_inject aligned_split_right append.assoc append_Cons append_Nil2 append_take_drop_id assms same_append_eq split_half.simps split_half_ts)
ultimately show ?thesis
using split_half_ts False by auto
qed simp
lemma node\<^sub>i_Laligned:
assumes "Laligned (Node ts t) u"
shows "Laligned_up\<^sub>i (node\<^sub>i k ts t) u"
using assms
proof (cases "length ts \<le> 2*k")
case False
then obtain ls sub sep rs where
split_half_ts: "split_half ts = (ls@[(sub,sep)], rs)"
by (meson node\<^sub>i_cases)
then have "Laligned (Node ls sub) sep"
by (metis Laligned_split_left append.assoc append_Cons assms split_half_conc)
moreover have "aligned sep (Node rs t) u"
by (metis Laligned_split_right append.assoc append_Cons append_Nil2 assms same_append_eq split_half_conc split_half_ts)
ultimately show ?thesis
using split_half_ts False by auto
qed simp
lemma length_right_side: "length xs > 1 \<Longrightarrow> length (drop ((length xs + 1) div 2) xs) > 0"
by auto
lemma Lnode\<^sub>i_aligned:
assumes "aligned l (Leaf ks) u"
and "sorted_less ks"
and "k > 0"
shows "aligned_up\<^sub>i l (Lnode\<^sub>i k ks) u"
using assms
proof (cases "length ks \<le> 2*k")
case False
then obtain ls sep rs where split_half_ts:
"take ((length ks + 1) div 2) ks = ls@[sep]"
"drop ((length ks + 1) div 2) ks = rs"
using split_half_not_empty[of ks]
by auto
moreover have "sorted_less (ls@[sep])"
by (metis append_take_drop_id assms(2) sorted_wrt_append split_half_ts(1))
ultimately have "aligned l (Leaf (ls@[sep])) sep"
using split_half_conc[of ks "ls@[sep]" rs] assms sorted_snoc_iff[of ls sep]
by auto
moreover have "aligned sep (Leaf rs) u"
proof -
have "length rs > 0"
using False assms(3) split_half_ts(2) by fastforce
then obtain sep' rs' where "rs = sep' # rs'"
by (cases rs) auto
moreover have "sep < sep'"
by (metis append_take_drop_id assms(2) calculation in_set_conv_decomp sorted_mid_iff sorted_snoc_iff split_half_ts(1) split_half_ts(2))
moreover have "sorted_less rs"
by (metis append_take_drop_id assms(2) sorted_wrt_append split_half_ts(2))
ultimately show ?thesis
using split_half_ts split_half_conc[of ks "ls@[sep]" rs] assms
by auto
qed
ultimately show ?thesis
using split_half_ts False by auto
qed simp
lemma height_up\<^sub>i_merge: "height_up\<^sub>i (Up\<^sub>i l a r) = height t \<Longrightarrow> height (Node (ls@(t,x)#rs) tt) = height (Node (ls@(l,a)#(r,x)#rs) tt)"
by simp
lemma ins_height: "height_up\<^sub>i (ins k x t) = height t"
proof(induction k x t rule: ins.induct)
case (2 k x ts t)
then obtain ls rs where split_list: "split ts x = (ls,rs)"
by (meson surj_pair)
then have split_append: "ts = ls@rs"
using split_conc
by auto
then show ?case
proof (cases rs)
case Nil
then have height_sub: "height_up\<^sub>i (ins k x t) = height t"
using 2 by (simp add: split_list)
then show ?thesis
proof (cases "ins k x t")
case (T\<^sub>i a)
then have "height (Node ts t) = height (Node ts a)"
using height_sub
by simp
then show ?thesis
using T\<^sub>i Nil split_list split_append
by simp
next
case (Up\<^sub>i l a r)
then have "height (Node ls t) = height (Node (ls@[(l,a)]) r)"
using height_bplustree_order height_sub by (induction ls) auto
then show ?thesis using 2 Nil split_list Up\<^sub>i split_append
by (simp del: node\<^sub>i.simps add: node\<^sub>i_height)
qed
next
case (Cons a list)
then obtain sub sep where a_split: "a = (sub,sep)"
by (cases a) auto
then have height_sub: "height_up\<^sub>i (ins k x sub) = height sub"
by (metis "2.IH"(2) a_split Cons split_list)
then show ?thesis
proof (cases "ins k x sub")
case (T\<^sub>i a)
then have "height a = height sub"
using height_sub by auto
then have "height (Node (ls@(sub,sep)#rs) t) = height (Node (ls@(a,sep)#rs) t)"
by auto
then show ?thesis
using T\<^sub>i height_sub Cons 2 split_list a_split split_append
by (auto simp add: image_Un max.commute finite_set_ins_swap)
next
case (Up\<^sub>i l a r)
then have "height (Node (ls@(sub,sep)#list) t) = height (Node (ls@(l,a)#(r,sep)#list) t)"
using height_up\<^sub>i_merge height_sub
by fastforce
then show ?thesis
using Up\<^sub>i Cons 2 split_list a_split split_append
by (auto simp del: node\<^sub>i.simps simp add: node\<^sub>i_height image_Un max.commute finite_set_ins_swap)
qed
qed
qed simp
(* the below proof is overly complicated as a number of lemmas regarding height are missing *)
lemma ins_bal: "bal t \<Longrightarrow> bal_up\<^sub>i (ins k x t)"
proof(induction k x t rule: ins.induct)
case (2 k x ts t)
then obtain ls rs where split_res: "split ts x = (ls, rs)"
by (meson surj_pair)
then have split_app: "ts = ls@rs"
using split_conc
by fastforce
show ?case
proof (cases rs)
case Nil
then show ?thesis
proof (cases "ins k x t")
case (T\<^sub>i a)
then have "bal (Node ls a)" unfolding bal.simps
by (metis "2.IH"(1) "2.prems" append_Nil2 bal.simps(2) bal_up\<^sub>i.simps(1) height_up\<^sub>i.simps(1) ins_height local.Nil split_app split_res)
then show ?thesis
using Nil T\<^sub>i 2 split_res
by simp
next
case (Up\<^sub>i l a r)
then have
"(\<forall>x\<in>set (subtrees (ls@[(l,a)])). bal x)"
"(\<forall>x\<in>set (subtrees ls). height r = height x)"
using 2 Up\<^sub>i Nil split_res split_app
by simp_all (metis height_up\<^sub>i.simps(2) ins_height max_def)
then show ?thesis unfolding ins.simps
using Up\<^sub>i Nil 2 split_res
by (simp del: node\<^sub>i.simps add: node\<^sub>i_bal)
qed
next
case (Cons a list)
then obtain sub sep where a_prod: "a = (sub, sep)" by (cases a)
then have "bal_up\<^sub>i (ins k x sub)" using 2 split_res
using a_prod local.Cons split_app by auto
show ?thesis
proof (cases "ins k x sub")
case (T\<^sub>i x1)
then have "height x1 = height t"
by (metis "2.prems" a_prod add_diff_cancel_left' bal_split_left(1) bal_split_left(2) height_bal_tree height_up\<^sub>i.simps(1) ins_height local.Cons plus_1_eq_Suc split_app)
then show ?thesis
using split_app Cons T\<^sub>i 2 split_res a_prod
by auto
next
case (Up\<^sub>i l a r)
(* The only case where explicit reasoning is required - likely due to the insertion of 2 elements in the list *)
then have
"\<forall>x \<in> set (subtrees (ls@(l,a)#(r,sep)#list)). bal x"
using Up\<^sub>i split_app Cons 2 \<open>bal_up\<^sub>i (ins k x sub)\<close> by auto
moreover have "\<forall>x \<in> set (subtrees (ls@(l,a)#(r,sep)#list)). height x = height t"
using Up\<^sub>i split_app Cons 2 \<open>bal_up\<^sub>i (ins k x sub)\<close> ins_height split_res a_prod
apply auto
by (metis height_up\<^sub>i.simps(2) sup.idem sup_nat_def)
ultimately show ?thesis using Up\<^sub>i Cons 2 split_res a_prod
by (simp del: node\<^sub>i.simps add: node\<^sub>i_bal)
qed
qed
qed simp
(* ins acts as ins_list wrt inorder *)
(* "simple enough" to be automatically solved *)
lemma node\<^sub>i_leaves: "leaves_up\<^sub>i (node\<^sub>i k ts t) = leaves (Node ts t)"
proof (cases "length ts \<le> 2*k")
case False
then obtain ls sub sep rs where
split_half_ts: "split_half ts = (ls@[(sub,sep)], rs)"
by (meson node\<^sub>i_cases)
then have "leaves_up\<^sub>i (node\<^sub>i k ts t) = leaves_list ls @ leaves sub @ leaves_list rs @ leaves t"
using False by auto
also have "\<dots> = leaves (Node ts t)"
using split_half_ts split_half_conc[of ts "ls@[(sub,sep)]" rs] by auto
finally show ?thesis.
qed simp
corollary node\<^sub>i_leaves_simps:
"node\<^sub>i k ts t = T\<^sub>i t' \<Longrightarrow> leaves t' = leaves (Node ts t)"
"node\<^sub>i k ts t = Up\<^sub>i l a r \<Longrightarrow> leaves l @ leaves r = leaves (Node ts t)"
apply (metis leaves_up\<^sub>i.simps(1) node\<^sub>i_leaves)
by (metis leaves_up\<^sub>i.simps(2) node\<^sub>i_leaves)
lemma Lnode\<^sub>i_leaves: "leaves_up\<^sub>i (Lnode\<^sub>i k xs) = leaves (Leaf xs)"
proof (cases "length xs \<le> 2*k")
case False
then obtain ls sub sep rs where
split_half_ts: "split_half xs = (ls@[sep], rs)"
by (meson Lnode\<^sub>i_cases)
then have "leaves_up\<^sub>i (Lnode\<^sub>i k xs) = ls @ sep # rs"
using False by auto
also have "\<dots> = leaves (Leaf xs)"
using split_half_ts split_half_conc[of xs "ls@[sep]" rs] by auto
finally show ?thesis.
qed simp
corollary Lnode\<^sub>i_leaves_simps:
"Lnode\<^sub>i k xs = T\<^sub>i t \<Longrightarrow> leaves t = leaves (Leaf xs)"
"Lnode\<^sub>i k xs = Up\<^sub>i l a r \<Longrightarrow> leaves l @ leaves r = leaves (Leaf xs)"
apply (metis leaves_up\<^sub>i.simps(1) Lnode\<^sub>i_leaves)
by (metis leaves_up\<^sub>i.simps(2) Lnode\<^sub>i_leaves)
(* specialize ins_list_sorted since it is cumbersome to express
"inorder_list ts" as "xs @ [a]" and always having to use the implicit properties of split*)
lemma ins_list_split:
assumes "Laligned (Node ts t) u"
and "sorted_less (leaves (Node ts t))"
and "split ts x = (ls, rs)"
shows "ins_list x (leaves (Node ts t)) = leaves_list ls @ ins_list x (leaves_list rs @ leaves t)"