Theory BTree_Split
theory BTree_Split
imports BTree_Set
begin
section "Abstract split functions"
subsection "Linear split"
text "Finally we show that the split axioms are feasible
by providing an example split function"
fun linear_split_help:: "(_×'a::linorder) list ⇒ _ ⇒ (_×_) list ⇒ ((_×_) list × (_×_) list)" where
"linear_split_help [] x prev = (prev, [])" |
"linear_split_help ((sub, sep)#xs) x prev = (if sep < x then linear_split_help xs x (prev @ [(sub, sep)]) else (prev, (sub,sep)#xs))"
fun linear_split:: "(_×'a::linorder) list ⇒ _ ⇒ ((_×_) list × (_×_) list)" where
"linear_split xs x = linear_split_help xs x []"
text "Linear split is similar to well known functions, therefore a quick proof can be done."
lemma linear_split_alt: "linear_split xs x = (takeWhile (λ(_,s). s<x) xs, dropWhile (λ(_,s). s<x) xs)"
proof -
have "linear_split_help xs x prev = (prev @ takeWhile (λ(_, s). s < x) xs, dropWhile (λ(_, s). s < x) xs)"
for prev
apply (induction xs arbitrary: prev)
apply auto
done
thus ?thesis by auto
qed