# HG changeset patch # User nipkow # Date 1503695396 -7200 # Node ID ca7a369301f6dc1abd97eec21659c417a3dd612a # Parent 65b6d48fc9a90aba13376f81033d84135a520fbd reorganization of tree lemmas; new lemmas diff -r 65b6d48fc9a9 -r ca7a369301f6 src/HOL/Data_Structures/Balance.thy --- a/src/HOL/Data_Structures/Balance.thy Fri Aug 25 13:01:13 2017 +0100 +++ b/src/HOL/Data_Structures/Balance.thy Fri Aug 25 23:09:56 2017 +0200 @@ -4,67 +4,9 @@ theory Balance imports - Complex_Main - "HOL-Library.Tree" + "HOL-Library.Tree_Real" begin -(* The following two lemmas should go into theory \Tree\, except that that -theory would then depend on \Complex_Main\. *) - -lemma min_height_balanced: assumes "balanced t" -shows "min_height t = nat(floor(log 2 (size1 t)))" -proof cases - assume *: "complete t" - hence "size1 t = 2 ^ min_height t" - by (simp add: complete_iff_height size1_if_complete) - hence "size1 t = 2 powr min_height t" - using * by (simp add: powr_realpow) - hence "min_height t = log 2 (size1 t)" - by simp - thus ?thesis - by linarith -next - assume *: "\ complete t" - hence "height t = min_height t + 1" - using assms min_height_le_height[of t] - by(auto simp add: balanced_def complete_iff_height) - hence "2 ^ min_height t \ size1 t \ size1 t < 2 ^ (min_height t + 1)" - by (metis * min_height_size1 size1_height_if_incomplete) - hence "2 powr min_height t \ size1 t \ size1 t < 2 powr (min_height t + 1)" - by(simp only: powr_realpow) - (metis of_nat_less_iff of_nat_le_iff of_nat_numeral of_nat_power) - hence "min_height t \ log 2 (size1 t) \ log 2 (size1 t) < min_height t + 1" - by(simp add: log_less_iff le_log_iff) - thus ?thesis by linarith -qed - -lemma height_balanced: assumes "balanced t" -shows "height t = nat(ceiling(log 2 (size1 t)))" -proof cases - assume *: "complete t" - hence "size1 t = 2 ^ height t" - by (simp add: size1_if_complete) - hence "size1 t = 2 powr height t" - using * by (simp add: powr_realpow) - hence "height t = log 2 (size1 t)" - by simp - thus ?thesis - by linarith -next - assume *: "\ complete t" - hence **: "height t = min_height t + 1" - using assms min_height_le_height[of t] - by(auto simp add: balanced_def complete_iff_height) - hence 0: "2 ^ min_height t < size1 t \ size1 t \ 2 ^ (min_height t + 1)" - by (metis "*" min_height_size1_if_incomplete size1_height) - hence "2 powr min_height t < size1 t \ size1 t \ 2 powr (min_height t + 1)" - by(simp only: powr_realpow) - (metis of_nat_less_iff of_nat_le_iff of_nat_numeral of_nat_power) - hence "min_height t < log 2 (size1 t) \ log 2 (size1 t) \ min_height t + 1" - by(simp add: log_le_iff less_log_iff) - thus ?thesis using ** by linarith -qed - (* mv *) text \The lemmas about \floor\ and \ceiling\ of \log 2\ should be generalized diff -r 65b6d48fc9a9 -r ca7a369301f6 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Fri Aug 25 13:01:13 2017 +0100 +++ b/src/HOL/Library/Library.thy Fri Aug 25 23:09:56 2017 +0200 @@ -80,6 +80,7 @@ Sum_of_Squares Transitive_Closure_Table Tree_Multiset + Tree_Real Type_Length While_Combinator begin diff -r 65b6d48fc9a9 -r ca7a369301f6 src/HOL/Library/Tree_Real.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Tree_Real.thy Fri Aug 25 23:09:56 2017 +0200 @@ -0,0 +1,64 @@ +(* Author: Tobias Nipkow *) + +theory Tree_Real +imports + Complex_Main + Tree +begin + +text \This theory is separate from @{theory Tree} because the former is discrete and builds on +@{theory Main} whereas this theory builds on @{theory Complex_Main}.\ + + +lemma size1_height_log: "log 2 (size1 t) \ height t" +by (simp add: log2_of_power_le size1_height) + +lemma min_height_size1_log: "min_height t \ log 2 (size1 t)" +by (simp add: le_log2_of_power min_height_size1) + +lemma size1_log_if_complete: "complete t \ height t = log 2 (size1 t)" +by (simp add: log2_of_power_eq size1_if_complete) + +lemma min_height_size1_log_if_incomplete: + "\ complete t \ min_height t < log 2 (size1 t)" +by (simp add: less_log2_of_power min_height_size1_if_incomplete) + + +lemma min_height_balanced: assumes "balanced t" +shows "min_height t = nat(floor(log 2 (size1 t)))" +proof cases + assume *: "complete t" + hence "size1 t = 2 ^ min_height t" + by (simp add: complete_iff_height size1_if_complete) + from log2_of_power_eq[OF this] show ?thesis by linarith +next + assume *: "\ complete t" + hence "height t = min_height t + 1" + using assms min_height_le_height[of t] + by(auto simp add: balanced_def complete_iff_height) + hence "size1 t < 2 ^ (min_height t + 1)" + by (metis * size1_height_if_incomplete) + hence "log 2 (size1 t) < min_height t + 1" + using log2_of_power_less size1_ge0 by blast + thus ?thesis using min_height_size1_log[of t] by linarith +qed + +lemma height_balanced: assumes "balanced t" +shows "height t = nat(ceiling(log 2 (size1 t)))" +proof cases + assume *: "complete t" + hence "size1 t = 2 ^ height t" + by (simp add: size1_if_complete) + from log2_of_power_eq[OF this] show ?thesis + by linarith +next + assume *: "\ complete t" + hence **: "height t = min_height t + 1" + using assms min_height_le_height[of t] + by(auto simp add: balanced_def complete_iff_height) + hence "size1 t \ 2 ^ (min_height t + 1)" by (metis size1_height) + from log2_of_power_le[OF this size1_ge0] min_height_size1_log_if_incomplete[OF *] ** + show ?thesis by linarith +qed + +end \ No newline at end of file diff -r 65b6d48fc9a9 -r ca7a369301f6 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Fri Aug 25 13:01:13 2017 +0100 +++ b/src/HOL/Transcendental.thy Fri Aug 25 23:09:56 2017 +0200 @@ -2772,21 +2772,66 @@ by (simp add: log_powr powr_realpow [symmetric]) lemma le_log_of_power: - assumes "1 < b" "b ^ n \ m" + assumes "b ^ n \ m" "1 < b" shows "n \ log b m" proof - - from assms have "0 < m" - by (metis less_trans zero_less_power less_le_trans zero_less_one) - have "n = log b (b ^ n)" - using assms(1) by (simp add: log_nat_power) - also have "\ \ log b m" - using assms \0 < m\ by simp + from assms have "0 < m" by (metis less_trans zero_less_power less_le_trans zero_less_one) + have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) + also have "\ \ log b m" using assms \0 < m\ by simp + finally show ?thesis . +qed + +lemma le_log2_of_power: "2 ^ n \ m \ n \ log 2 m" for m n :: nat +using le_log_of_power[of 2] by simp + +lemma log_of_power_le: + assumes "m \ b ^ n" "b > 1" "m > 0" + shows "log b (real m) \ n" +proof - + have "log b m \ log b (b ^ n)" using assms by simp + also have "\ = n" using assms(2) by (simp add: log_nat_power) + finally show ?thesis . +qed + +lemma log2_of_power_le: "\ m \ 2 ^ n; m > 0 \ \ log 2 m \ n" for m n :: nat +using log_of_power_le[of _ 2] by simp + +lemma log_of_power_less: + assumes "m < b ^ n" "b > 1" "m > 0" + shows "log b (real m) < n" +proof - + have "log b m < log b (b ^ n)" using assms by simp + also have "\ = n" using assms(2) by (simp add: log_nat_power) + finally show ?thesis . +qed + +lemma log2_of_power_less: "\ m < 2 ^ n; m > 0 \ \ log 2 m < n" for m n :: nat +using log_of_power_less[of _ 2] by simp + +lemma log_of_power_eq: + assumes "m = b ^ n" "b > 1" + shows "n = log b (real m)" +proof - + have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) + also have "\ = log b m" using assms by (simp) + finally show ?thesis . +qed + +lemma log2_of_power_eq: "m = 2 ^ n \ n = log 2 m" for m n :: nat +using log_of_power_eq[of _ 2] by simp + +lemma less_log_of_power: + assumes "b ^ n < m" "1 < b" + shows "n < log b m" +proof - + have "0 < m" by (metis assms less_trans zero_less_power zero_less_one) + have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) + also have "\ < log b m" using assms \0 < m\ by simp finally show ?thesis . qed -lemma le_log2_of_power: "2 ^ n \ m \ n \ log 2 m" - for m n :: nat - using le_log_of_power[of 2] by simp +lemma less_log2_of_power: "2 ^ n < m \ n < log 2 m" for m n :: nat +using less_log_of_power[of 2] by simp lemma log_base_change: "0 < a \ a \ 1 \ log b x = log a x / log a b" by (simp add: log_def)