Theory KD_Tree
section "Definition of the ‹k›-d Tree"
theory KD_Tree
imports
Complex_Main
"HOL-Analysis.Finite_Cartesian_Product"
"HOL-Analysis.Topology_Euclidean_Space"
begin
text ‹
A ‹k›-d tree is a space-partitioning data structure for organizing points in a ‹k›-dimensional space.
In principle the ‹k›-d tree is a binary tree. The leafs hold the ‹k›-dimensional points and the nodes
contain left and right subtrees as well as a discriminator ‹v› at a particular axis ‹k›.
Every node divides the space into two parts by splitting along a hyperplane.
Consider a node ‹n› with associated discriminator ‹v› at axis ‹k›.
All points in the left subtree must have a value at axis ‹k› that is less than or
equal to ‹v› and all points in the right subtree must have a value at axis ‹k› that is
greater than ‹v›.
Deviations from the papers:
The chosen tree representation is taken from \<^cite>‹"DBLP:journals/toms/FriedmanBF77"› with one minor
adjustment. Originally the leafs hold buckets of points of size ‹b›. This representation fixes the
bucket size to ‹b = 1›, a single point per Leaf. This is only a minor adjustment since the paper
proves that ‹b = 1› is the optimal bucket size for minimizing the running time of the nearest neighbor
algorithm \<^cite>‹"DBLP:journals/toms/FriedmanBF77"›, only simplifies building the optimized
‹k›-d trees \<^cite>‹"DBLP:journals/toms/FriedmanBF77"› and has little influence on the
search algorithm \<^cite>‹"DBLP:journals/cacm/Bentley75"›.
›
type_synonym 'k point = "(real, 'k) vec"
lemma dist_point_def:
fixes p⇩0 :: "('k::finite) point"
shows "dist p⇩0 p⇩1 = sqrt (∑k ∈ UNIV. (p⇩0$k - p⇩1$k)⇧2)"
unfolding dist_vec_def L2_set_def dist_real_def by simp