# HG changeset patch # User nipkow # Date 1721311240 -7200 # Node ID 27e66a8323b2de7de5b6f525ee9670d4adf9c651 # Parent 01edf83f6dee088595ac6365a452df5ad6a973a0# Parent 97dab09aa727ec2cfefc39f8f2a7b5204a8eadb6 merged diff -r 01edf83f6dee -r 27e66a8323b2 src/HOL/Data_Structures/AVL_Bal_Set.thy --- a/src/HOL/Data_Structures/AVL_Bal_Set.thy Thu Jul 18 13:52:51 2024 +0200 +++ b/src/HOL/Data_Structures/AVL_Bal_Set.thy Thu Jul 18 16:00:40 2024 +0200 @@ -67,13 +67,13 @@ GT \ let r' = insert x r in if incr r r' then balR l a b r' else Node l (a,b) r')" fun decr where -"decr t t' = (t \ Leaf \ (t' = Leaf \ \ is_bal t \ is_bal t'))" +"decr t t' = (t \ Leaf \ incr t' t)" fun split_max :: "'a tree_bal \ 'a tree_bal * 'a" where "split_max (Node l (a, ba) r) = (if r = Leaf then (l,a) else let (r',a') = split_max r; - t' = if decr r r' then balL l a ba r' else Node l (a,ba) r' + t' = if incr r' r then balL l a ba r' else Node l (a,ba) r' in (t', a'))" fun delete :: "'a::linorder \ 'a tree_bal \ 'a tree_bal" where @@ -82,7 +82,7 @@ (case cmp x a of EQ \ if l = Leaf then r else let (l', a') = split_max l in - if decr l l' then balR l' a' ba r else Node l' (a',ba) r | + if incr l' l then balR l' a' ba r else Node l' (a',ba) r | LT \ let l' = delete x l in if decr l l' then balR l' a ba r else Node l' (a,ba) r | GT \ let r' = delete x r in if decr r r' then balL l a ba r' else Node l (a,ba) r')" @@ -136,7 +136,7 @@ lemma avl_split_max: "\ split_max t = (t',a); avl t; t \ Leaf \ \ - avl t' \ height t = height t' + (if decr t t' then 1 else 0)" + avl t' \ height t = height t' + (if incr t' t then 1 else 0)" apply(induction t arbitrary: t' a rule: split_max_induct) apply(auto simp: max_absorb1 max_absorb2 height_1_iff split!: splits prod.splits) done diff -r 01edf83f6dee -r 27e66a8323b2 src/HOL/Data_Structures/RBT_Set.thy --- a/src/HOL/Data_Structures/RBT_Set.thy Thu Jul 18 13:52:51 2024 +0200 +++ b/src/HOL/Data_Structures/RBT_Set.thy Thu Jul 18 16:00:40 2024 +0200 @@ -307,10 +307,13 @@ lemma bheight_size_bound: "invc t \ invh t \ 2 ^ (bheight t) \ size1 t" by (induction t) auto +lemma bheight_le_min_height: "invh t \ bheight t \ min_height t" +by (induction t) auto + lemma rbt_height_le: assumes "rbt t" shows "height t \ 2 * log 2 (size1 t)" proof - have "2 powr (height t / 2) \ 2 powr bheight t" - using rbt_height_bheight[OF assms] by (simp) + using rbt_height_bheight[OF assms] by simp also have "\ \ size1 t" using assms by (simp add: powr_realpow bheight_size_bound rbt_def) finally have "2 powr (height t / 2) \ size1 t" . @@ -319,4 +322,15 @@ thus ?thesis by simp qed +lemma rbt_height_le2: assumes "rbt t" shows "height t \ 2 * log 2 (size1 t)" +proof - + have "height t \ 2 * bheight t" + using rbt_height_bheight_if assms[simplified rbt_def] by fastforce + also have "\ \ 2 * min_height t" + using bheight_le_min_height assms[simplified rbt_def] by auto + also have "\ \ 2 * log 2 (size1 t)" + using le_log2_of_power min_height_size1 by auto + finally show ?thesis by simp +qed + end