diff -r ed5b907bbf50 -r 831f17da1aab src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Sun Nov 08 21:27:08 2020 +0100 +++ b/src/HOL/Library/Tree.thy Tue Nov 10 17:42:41 2020 +0100 @@ -1,5 +1,5 @@ (* Author: Tobias Nipkow *) -(* Todo: minimal ipl of balanced trees *) +(* Todo: minimal ipl of almost complete trees *) section \Binary Tree\ @@ -55,8 +55,9 @@ "complete Leaf = True" | "complete (Node l x r) = (complete l \ complete r \ height l = height r)" -definition balanced :: "'a tree \ bool" where -"balanced t = (height t - min_height t \ 1)" +text \Almost complete:\ +definition acomplete :: "'a tree \ bool" where +"acomplete t = (height t - min_height t \ 1)" text \Weight balanced:\ fun wbalanced :: "'a tree \ bool" where @@ -290,24 +291,24 @@ using complete_if_size1_height size1_if_complete by blast -subsection \\<^const>\balanced\\ +subsection \\<^const>\acomplete\\ -lemma balanced_subtreeL: "balanced (Node l x r) \ balanced l" -by(simp add: balanced_def) +lemma acomplete_subtreeL: "acomplete (Node l x r) \ acomplete l" +by(simp add: acomplete_def) -lemma balanced_subtreeR: "balanced (Node l x r) \ balanced r" -by(simp add: balanced_def) +lemma acomplete_subtreeR: "acomplete (Node l x r) \ acomplete r" +by(simp add: acomplete_def) -lemma balanced_subtrees: "\ balanced t; s \ subtrees t \ \ balanced s" +lemma acomplete_subtrees: "\ acomplete t; s \ subtrees t \ \ acomplete s" using [[simp_depth_limit=1]] by(induction t arbitrary: s) - (auto simp add: balanced_subtreeL balanced_subtreeR) + (auto simp add: acomplete_subtreeL acomplete_subtreeR) text\Balanced trees have optimal height:\ -lemma balanced_optimal: +lemma acomplete_optimal: fixes t :: "'a tree" and t' :: "'b tree" -assumes "balanced t" "size t \ size t'" shows "height t \ height t'" +assumes "acomplete t" "size t \ size t'" shows "height t \ height t'" proof (cases "complete t") case True have "(2::nat) ^ height t \ 2 ^ height t'" @@ -333,7 +334,7 @@ hence *: "min_height t < height t'" by simp have "min_height t + 1 = height t" using min_height_le_height[of t] assms(1) False - by (simp add: complete_iff_height balanced_def) + by (simp add: complete_iff_height acomplete_def) with * show ?thesis by arith qed