876 qualified lemma compute_float_up[code]: "float_up p x = - float_down p (-x)" |
876 qualified lemma compute_float_up[code]: "float_up p x = - float_down p (-x)" |
877 by transfer (simp add: round_down_uminus_eq) |
877 by transfer (simp add: round_down_uminus_eq) |
878 |
878 |
879 end |
879 end |
880 |
880 |
881 definition floorlog :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
882 where "floorlog b a = (if a > 0 \<and> b > 1 then nat \<lfloor>log b a\<rfloor> + 1 else 0)" |
|
883 |
|
884 lemma floorlog_mono: "x \<le> y \<Longrightarrow> floorlog b x \<le> floorlog b y" |
|
885 by(auto simp: floorlog_def floor_mono nat_mono) |
|
886 |
|
887 lemma floorlog_bounds: |
|
888 assumes "x > 0" "b > 1" |
|
889 shows "b ^ (floorlog b x - 1) \<le> x \<and> x < b ^ (floorlog b x)" |
|
890 proof |
|
891 show "b ^ (floorlog b x - 1) \<le> x" |
|
892 proof - |
|
893 have "(b::real) ^ nat \<lfloor>log b (real_of_int x)\<rfloor> = b powr real_of_int \<lfloor>log b (real_of_int x)\<rfloor>" |
|
894 using powr_realpow[symmetric, of b "nat \<lfloor>log b (real_of_int x)\<rfloor>"] \<open>x > 0\<close> \<open>b > 1\<close> |
|
895 by simp |
|
896 also have "\<dots> \<le> b powr log b (real_of_int x)" |
|
897 using \<open>b > 1\<close> |
|
898 by simp |
|
899 also have "\<dots> = real_of_int x" |
|
900 using \<open>0 < x\<close> \<open>b > 1\<close> by simp |
|
901 finally have "b ^ nat \<lfloor>log b (real_of_int x)\<rfloor> \<le> real_of_int x" |
|
902 by simp |
|
903 then show ?thesis |
|
904 using \<open>0 < x\<close> \<open>b > 1\<close> of_nat_le_iff |
|
905 by (fastforce simp add: floorlog_def) |
|
906 qed |
|
907 show "x < b ^ (floorlog b x)" |
|
908 proof - |
|
909 have "x \<le> b powr (log b x)" |
|
910 using \<open>x > 0\<close> \<open>b > 1\<close> by simp |
|
911 also have "\<dots> < b powr (\<lfloor>log b x\<rfloor> + 1)" |
|
912 using assms |
|
913 by (intro powr_less_mono) auto |
|
914 also have "\<dots> = b ^ nat (\<lfloor>log b (real_of_int x)\<rfloor> + 1)" |
|
915 using assms |
|
916 by (simp add: powr_realpow[symmetric]) |
|
917 finally |
|
918 have "x < b ^ nat (\<lfloor>log (real b) (real_of_int (int x))\<rfloor> + 1)" |
|
919 by (rule of_nat_less_imp_less) |
|
920 then show ?thesis |
|
921 using \<open>x > 0\<close> \<open>b > 1\<close> |
|
922 by (simp add: floorlog_def nat_add_distrib) |
|
923 qed |
|
924 qed |
|
925 |
|
926 lemma floorlog_power[simp]: |
|
927 assumes "a > 0" "b > 1" |
|
928 shows "floorlog b (a * b ^ c) = floorlog b a + c" |
|
929 proof - |
|
930 have "\<lfloor>log (real b) (real a) + real c\<rfloor> = \<lfloor>log (real b) (real a)\<rfloor> + c" |
|
931 by arith |
|
932 then show ?thesis using assms |
|
933 by (auto simp: floorlog_def log_mult powr_realpow[symmetric] nat_add_distrib) |
|
934 qed |
|
935 |
|
936 lemma floor_log_add_eqI: |
|
937 fixes a::nat and b::nat and r::real |
|
938 assumes "b > 1" "a \<ge> 1" "0 \<le> r" "r < 1" |
|
939 shows "\<lfloor>log b (a + r)\<rfloor> = \<lfloor>log b a\<rfloor>" |
|
940 proof (rule floor_eq2) |
|
941 have "log (real b) (real a) \<le> log (real b) (real a + r)" |
|
942 using assms by force |
|
943 then show "\<lfloor>log (real b) (real a)\<rfloor> \<le> log (real b) (real a + r)" |
|
944 by arith |
|
945 next |
|
946 define l::int where "l = int b ^ (nat \<lfloor>log b a\<rfloor> + 1)" |
|
947 have l_def_real: "l = b powr (\<lfloor>log b a\<rfloor> + 1)" |
|
948 using assms by (simp add: l_def powr_add powr_real_of_int) |
|
949 have "a < l" |
|
950 proof - |
|
951 have "a = b powr (log b a)" using assms by simp |
|
952 also have "\<dots> < b powr floor ((log b a) + 1)" |
|
953 using assms(1) by auto |
|
954 also have "\<dots> = l" |
|
955 using assms |
|
956 by (simp add: l_def powr_real_of_int powr_add) |
|
957 finally show ?thesis by simp |
|
958 qed |
|
959 then have "a + r < l" using assms by simp |
|
960 then have "log b (a + r) < log b l" |
|
961 using assms by simp |
|
962 also have "\<dots> = real_of_int \<lfloor>log b a\<rfloor> + 1" |
|
963 using assms by (simp add: l_def_real) |
|
964 finally show "log b (a + r) < real_of_int \<lfloor>log b a\<rfloor> + 1" . |
|
965 qed |
|
966 |
|
967 lemma divide_nat_diff_div_nat_less_one: |
|
968 fixes x b::nat shows "x / b - x div b < 1" |
|
969 proof - |
|
970 have "int 0 \<noteq> \<lfloor>1::real\<rfloor>" by simp |
|
971 thus ?thesis |
|
972 by (metis add_diff_cancel_left' floor_divide_of_nat_eq less_eq_real_def |
|
973 mod_div_trivial real_of_nat_div3 real_of_nat_div_aux) |
|
974 qed |
|
975 |
|
976 lemma floor_log_div: |
|
977 fixes b x :: nat assumes "b > 1" "x > 0" "x div b > 0" |
|
978 shows "\<lfloor>log b x\<rfloor> = \<lfloor>log b (x div b)\<rfloor> + 1" |
|
979 proof- |
|
980 have "\<lfloor>log (real b) (real x)\<rfloor> = \<lfloor>log (real b) (x / b * b)\<rfloor>" |
|
981 using assms by simp |
|
982 also have "\<dots> = \<lfloor>log b (x / b) + log b b\<rfloor>" |
|
983 using assms by (subst log_mult) auto |
|
984 also have "\<dots> = \<lfloor>log b (x / b)\<rfloor> + 1" using assms by simp |
|
985 also have "\<lfloor>log b (x / b)\<rfloor> = \<lfloor>log b (x div b + (x / b - x div b))\<rfloor>" |
|
986 by simp |
|
987 also have "\<dots> = \<lfloor>log b (x div b)\<rfloor>" |
|
988 using assms real_of_nat_div4 divide_nat_diff_div_nat_less_one |
|
989 by (intro floor_log_add_eqI) auto |
|
990 finally show ?thesis . |
|
991 qed |
|
992 |
|
993 lemma compute_floorlog[code]: |
|
994 "floorlog b x = (if x > 0 \<and> b > 1 then floorlog b (x div b) + 1 else 0)" |
|
995 by (auto simp: floorlog_def floor_log_div[of b x] div_eq_0_iff nat_add_distrib |
|
996 intro!: floor_eq2) |
|
997 |
|
998 lemma floor_log_eq_if: |
|
999 fixes b x y :: nat |
|
1000 assumes "x div b = y div b" "b > 1" "x > 0" "x div b \<ge> 1" |
|
1001 shows "floor(log b x) = floor(log b y)" |
|
1002 proof - |
|
1003 have "y > 0" using assms by(auto intro: ccontr) |
|
1004 thus ?thesis using assms by (simp add: floor_log_div) |
|
1005 qed |
|
1006 |
|
1007 lemma floorlog_eq_if: |
|
1008 fixes b x y :: nat |
|
1009 assumes "x div b = y div b" "b > 1" "x > 0" "x div b \<ge> 1" |
|
1010 shows "floorlog b x = floorlog b y" |
|
1011 proof - |
|
1012 have "y > 0" using assms by(auto intro: ccontr) |
|
1013 thus ?thesis using assms |
|
1014 by(auto simp add: floorlog_def eq_nat_nat_iff intro: floor_log_eq_if) |
|
1015 qed |
|
1016 |
|
1017 |
|
1018 definition bitlen :: "int \<Rightarrow> int" where "bitlen a = floorlog 2 (nat a)" |
|
1019 |
|
1020 lemma bitlen_alt_def: "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)" |
|
1021 by (simp add: bitlen_def floorlog_def) |
|
1022 |
|
1023 lemma bitlen_nonneg: "0 \<le> bitlen x" |
|
1024 by (simp add: bitlen_def) |
|
1025 |
|
1026 lemma bitlen_bounds: |
|
1027 assumes "x > 0" |
|
1028 shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)" |
|
1029 proof - |
|
1030 from assms have "bitlen x \<ge> 1" by (auto simp: bitlen_alt_def) |
|
1031 with assms floorlog_bounds[of "nat x" 2] show ?thesis |
|
1032 by (auto simp add: bitlen_def le_nat_iff nat_less_iff nat_diff_distrib) |
|
1033 qed |
|
1034 |
|
1035 lemma bitlen_pow2[simp]: |
|
1036 assumes "b > 0" |
|
1037 shows "bitlen (b * 2 ^ c) = bitlen b + c" |
|
1038 using assms |
|
1039 by (simp add: bitlen_def nat_mult_distrib nat_power_eq) |
|
1040 |
|
1041 lemma bitlen_Float: |
|
1042 fixes m e |
|
1043 defines "f \<equiv> Float m e" |
|
1044 shows "bitlen \<bar>mantissa f\<bar> + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)" |
|
1045 proof (cases "m = 0") |
|
1046 case True |
|
1047 then show ?thesis by (simp add: f_def bitlen_alt_def Float_def) |
|
1048 next |
|
1049 case False |
|
1050 then have "f \<noteq> float_of 0" |
|
1051 unfolding real_of_float_eq by (simp add: f_def) |
|
1052 then have "mantissa f \<noteq> 0" |
|
1053 by (simp add: mantissa_noteq_0) |
|
1054 moreover |
|
1055 obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i" |
|
1056 by (rule f_def[THEN denormalize_shift, OF \<open>f \<noteq> float_of 0\<close>]) |
|
1057 ultimately show ?thesis by (simp add: abs_mult) |
|
1058 qed |
|
1059 |
|
1060 context |
|
1061 begin |
|
1062 |
|
1063 qualified lemma compute_bitlen[code]: "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)" |
|
1064 unfolding bitlen_def |
|
1065 by (subst Float.compute_floorlog) (simp add: nat_div_distrib) |
|
1066 |
|
1067 end |
|
1068 |
881 |
1069 lemma float_gt1_scale: |
882 lemma float_gt1_scale: |
1070 assumes "1 \<le> Float m e" |
883 assumes "1 \<le> Float m e" |
1071 shows "0 \<le> e + (bitlen m - 1)" |
884 shows "0 \<le> e + (bitlen m - 1)" |
1072 proof - |
885 proof - |