diff --git a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy --- a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy +++ b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy @@ -1,349 +1,346 @@ (* Author: Fabian Hellauer Fabian Immler *) theory Conversion_IEEE_Float imports "HOL-Library.Float" IEEE_Properties "HOL-Library.Code_Target_Numeral" begin definition "of_finite (x::('e, 'f)float) = (if is_normal x then (Float (normal_mantissa x) (normal_exponent x)) else if is_denormal x then (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))) else 0)" lemma float_val_of_finite: "is_finite x \ of_finite x = valof x" by (induction x) (auto simp: normal_imp_not_denormal of_finite_def) definition is_normal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_normal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ fracwidth x + 1 \ - int (bias x) - bitlen \mantissa f\ + 1 < Float.exponent f \ Float.exponent f < 2^(LENGTH('e)) - bitlen \mantissa f\ - bias x" definition is_denormal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_denormal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias x) \ 1 - 2^(LENGTH('e) - 1) - int LENGTH('f) < Float.exponent f" lemmas is_denormal_FloatD = is_denormal_Float_def[THEN iffD1, THEN conjunct1] is_denormal_Float_def[THEN iffD1, THEN conjunct2] definition is_finite_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_finite_Float x f \ is_normal_Float x f \ is_denormal_Float x f \ f = 0" lemma is_finite_Float_eq: "is_finite_Float TYPE(('e, 'f)float) f \ (let e = Float.exponent f; bm = bitlen (abs (mantissa f)) in bm \ Suc LENGTH('f) \ bm \ 2 ^ (LENGTH('e) - 1) - e \ 1 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) < e)" proof - have *: "(2::int) ^ (LENGTH('e) - Suc 0) - 1 < 2 ^ LENGTH('e)" by (metis Suc_1 diff_le_self lessI linorder_not_less one_less_numeral_iff power_strict_increasing_iff zle_diff1_eq) have **: "1 - 2 ^ (LENGTH('e) - Suc 0) < int LENGTH('f)" by (smt len_gt_0 of_nat_0_less_iff zero_less_power) have ***: "2 ^ (LENGTH('e) - 1) + 1 = 2 ^ LENGTH('e) - int (bias TYPE(('e, 'f) IEEE.float))" by (simp add: bias_def power_Suc[symmetric] of_nat_diff) have rewr: "x \ 2 ^ n - e \ x + e < 2 ^ n + 1" for x::int and n e by auto show ?thesis unfolding *** rewr using * ** unfolding is_finite_Float_def is_normal_Float_def is_denormal_Float_def by (auto simp: Let_def bias_def mantissa_eq_zero_iff of_nat_diff intro: le_less_trans[OF add_right_mono]) qed lift_definition normal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m > 0 then 0 else 1, word_of_int (e + int (bias TYPE(('e, 'f)float)) + bitlen \m\ - 1), word_of_int (\m\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \m\)) - 2 ^ (LENGTH('f))))" . lemma sign_normal_of_Float:"sign (normal_of_Float x) = (if x > 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_pos_iff) -lemma uints_bitlen_eq: "uints n = {i. 0 \ i \ bitlen i \ n}" - by (auto simp: uints_num bitlen_le_iff_power) - lemma uint_word_of_int_bitlen_eq: "uint (word_of_int x::'a::len word) = x" if "bitlen x \ LENGTH('a)" "x \ 0" - by (subst word_uint.Abs_inverse) (simp_all add: uints_bitlen_eq that) + using that by (simp add: bitlen_le_iff_power take_bit_int_eq_self) lemma fraction_normal_of_Float:"fraction (normal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f))" if "is_normal_Float TYPE(('e, 'f)float) x" proof - from that have bmp: "bitlen \mantissa x\ > 0" by (metis abs_of_nonneg bitlen_bounds bitlen_def is_normal_Float_def nat_code(2) of_nat_0_le_iff power.simps(1) zabs_less_one_iff zero_less_abs_iff) have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have lem: "2 ^ nat (bitlen \mantissa x\ - 1) \ \mantissa x\" using bitlen_bounds is_normal_Float_def that zero_less_abs_iff by blast from that have nble: "nat (bitlen \mantissa x\) \ Suc LENGTH('f)" using bitlen_bounds by (auto simp: is_normal_Float_def) have nn: "0 \ \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)" apply (rule add_le_imp_le_diff) apply (rule order_trans[rotated]) apply (rule mult_right_mono) apply (rule lem, force) unfolding power_add[symmetric] using nble bmp by (auto) have "\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) < 2 * 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless) apply force unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) using nble by auto then have "bitlen (\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) \ int LENGTH('f)" unfolding bitlen_le_iff_power by simp then show ?thesis apply (transfer fixing: x) unfolding Let_def split_beta' fst_conv snd_conv uint_nat [symmetric] nat_uint_eq [symmetric] using nn apply (subst uint_word_of_int_bitlen_eq) apply (auto simp: nat_mult_distrib nat_diff_distrib nat_power_eq) done qed lemma exponent_normal_of_Float:"exponent (normal_of_Float x::('e, 'f)float) = nat (Float.exponent x + (bias TYPE(('e, 'f)float)) + bitlen \mantissa x\ - 1)" if "is_normal_Float TYPE(('e, 'f)float) x" using that apply (transfer fixing: x) apply (simp flip: uint_nat nat_uint_eq add: Let_def) apply (auto simp: is_normal_Float_def bitlen_le_iff_power uint_word_of_int_bitlen_eq Let_def) apply transfer apply (simp add: nat_take_bit_eq take_bit_int_eq_self) done lift_definition denormal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m \ 0 then 0 else 1, 0, word_of_int (\m\ * 2 ^ nat (e + bias TYPE(('e, 'f)float) + fracwidth TYPE(('e, 'f)float) - 1)))" . lemma sign_denormal_of_Float:"sign (denormal_of_Float x) = (if x \ 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_nonneg_iff) lemma exponent_denormal_of_Float:"exponent (denormal_of_Float x::('e, 'f)float) = 0" by (transfer fixing: x) (auto simp: Let_def) lemma fraction_denormal_of_Float:"fraction (denormal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ nat (Float.exponent x + bias TYPE(('e, 'f)float) + LENGTH('f) - 1))" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have *: "nat (bitlen \mantissa x\) + nat (Float.exponent x + (2 ^ (LENGTH('e) - Suc 0) + int LENGTH('f)) - 2) \ LENGTH('f)" using that by (auto simp: is_denormal_Float_def nat_diff_distrib' le_diff_conv bitlen_nonneg nat_le_iff bias_def nat_add_distrib[symmetric] of_nat_diff) have "\mantissa x\ * 2 ^ nat (Float.exponent x + int (bias TYPE(('e, 'f)float)) + LENGTH('f) - 1) < 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless, force) unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) apply (auto simp: bias_def) using that * by (auto simp: is_denormal_Float_def algebra_simps of_nat_diff) then show ?thesis apply (transfer fixing: x) apply transfer apply (simp add: Let_def nat_eq_iff take_bit_eq_mod) done qed definition of_finite_Float :: "Float.float \ ('e, 'f) float" where "of_finite_Float x = (if is_normal_Float TYPE(('e, 'f)float) x then normal_of_Float x else if is_denormal_Float TYPE(('e, 'f)float) x then denormal_of_Float x else 0)" lemma valof_normal_of_Float: "valof (normal_of_Float x::('e, 'f)float) = x" if "is_normal_Float TYPE(('e, 'f)float) x" proof - have "valof (normal_of_Float x::('e, 'f)float) = (- 1) ^ sign (normal_of_Float x::('e, 'f)float) * ((1 + real (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) / 2 ^ LENGTH('f)) * 2 powr (bitlen \mantissa x\ - 1)) * 2 powr Float.exponent x" (is "_ = ?s * ?m * ?e") using that by (auto simp: is_normal_Float_def valof_eq fraction_normal_of_Float powr_realpow[symmetric] exponent_normal_of_Float powr_diff powr_add) also have "\mantissa x\ > 0" using that by (auto simp: is_normal_Float_def) have bound: "2 ^ LENGTH('f) \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" proof - have "(2::nat) ^ LENGTH('f) \ 2 ^ nat (bitlen \mantissa x\ - 1) * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" by (simp add: power_add[symmetric]) also have "\ \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" using bitlen_bounds[of "\mantissa x\"] that by (auto simp: is_normal_Float_def) finally show ?thesis . qed have "?m = abs (mantissa x)" apply (subst of_nat_diff) subgoal using bound by auto subgoal using that by (auto simp: powr_realpow[symmetric] powr_add[symmetric] is_normal_Float_def bitlen_nonneg of_nat_diff divide_simps) done finally show ?thesis by (auto simp: mantissa_exponent sign_normal_of_Float abs_real_def zero_less_mult_iff) qed lemma valof_denormal_of_Float: "valof (denormal_of_Float x::('e, 'f)float) = x" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have less: "0 < Float.exponent x + (int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f))" using that by (auto simp: is_denormal_Float_def bias_def of_nat_diff) have "valof (denormal_of_Float x::('e, 'f)float) = ((- 1) ^ sign (denormal_of_Float x::('e, 'f)float) * \real_of_int (mantissa x)\) * (2 powr real (nat (Float.exponent x + int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f) - 1)) / (2 powr real (bias TYPE(('e, 'f) IEEE.float)) * 2 powr LENGTH('f)) * 2)" (is "_ = ?m * ?e") by (auto simp: valof_eq exponent_denormal_of_Float fraction_denormal_of_Float that mantissa_exponent powr_realpow[symmetric]) also have "?m = mantissa x" by (auto simp: sign_denormal_of_Float abs_real_def mantissa_neg_iff) also have "?e = 2 powr Float.exponent x" by (auto simp: powr_add[symmetric] divide_simps powr_mult_base less ac_simps) finally show ?thesis by (simp add: mantissa_exponent) qed lemma valof_of_finite_Float: "is_finite_Float (TYPE(('e, 'f) IEEE.float)) x \ valof (of_finite_Float x::('e, 'f)float) = x" by (auto simp: of_finite_Float_def is_finite_Float_def valof_denormal_of_Float valof_normal_of_Float) lemma is_normal_normal_of_Float: "is_normal (normal_of_Float x::('e, 'f)float)" if "is_normal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_normal_def exponent_normal_of_Float that is_normal_Float_def emax_eq nat_less_iff of_nat_diff) lemma is_denormal_denormal_of_Float: "is_denormal (denormal_of_Float x::('e, 'f)float)" if "is_denormal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_denormal_def exponent_denormal_of_Float that is_denormal_Float_def emax_eq fraction_denormal_of_Float le_nat_iff bias_def) lemma is_finite_of_finite_Float: "is_finite (of_finite_Float x)" by (auto simp: is_finite_def of_finite_Float_def is_normal_normal_of_Float is_denormal_denormal_of_Float) lemma Float_eq_zero_iff: "Float m e = 0 \ m = 0" by (metis Float.compute_is_float_zero Float_0_eq_0) lemma bitlen_mantissa_Float: shows "bitlen \mantissa (Float m e)\ = (if m = 0 then 0 else bitlen \m\ + e) - Float.exponent (Float m e)" using bitlen_Float[of m e] by auto lemma exponent_Float: shows "Float.exponent (Float m e) = (if m = 0 then 0 else bitlen \m\ + e) - bitlen \mantissa (Float m e)\ " using bitlen_Float[of m e] by auto lemma is_normal_Float_normal: "is_normal_Float TYPE(('e, 'f)float) (Float (normal_mantissa x) (normal_exponent x))" if "is_normal x" for x::"('e, 'f)float" proof - define f where "f = Float (normal_mantissa x) (normal_exponent x)" from that have "f \ 0" by (auto simp: f_def is_normal_def zero_float_def[symmetric] Float_eq_zero_iff normal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "normal_mantissa x = mantissa f * 2 ^ i" "normal_exponent x = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "normal_exponent x \ Float.exponent f" unfolding i by simp then have " bitlen \mantissa f\ \ 1 + int LENGTH('f)" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by auto moreover have "- int (bias TYPE(('e, 'f)float)) - bitlen \mantissa f\ + 1 < Float.exponent f" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def using that by (auto simp: mantissa_eq_zero_iff abs_mult bias_def normal_mantissa_def normal_exponent_def is_normal_def emax_eq less_diff_conv add_nonneg_eq_0_iff) moreover have "2 ^ (LENGTH('e) - Suc 0) + - (1::int) * 2 ^ LENGTH('e) \ 0" by simp then have "(2::int) ^ (LENGTH('e) - Suc 0) < 1 + 2 ^ LENGTH('e)" by arith then have "Float.exponent f < 2 ^ LENGTH('e) - bitlen \mantissa f\ - int (bias TYPE(('e, 'f)float))" using normal_exponent_bounds_int[OF that] unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by (auto simp: bias_def algebra_simps power_Suc[symmetric] of_nat_diff intro: le_less_trans[OF add_right_mono] normal_exponent_bounds_int[OF that]) ultimately show ?thesis by (auto simp: is_normal_Float_def f_def) qed lemma is_denormal_Float_denormal: "is_denormal_Float TYPE(('e, 'f)float) (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float)))" if "is_denormal x" for x::"('e, 'f)float" proof - define f where "f = Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))" from that have "f \ 0" by (auto simp: f_def is_denormal_def zero_float_def[symmetric] Float_eq_zero_iff denormal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "denormal_mantissa x = mantissa f * 2 ^ i" "denormal_exponent TYPE(('e, 'f)float) = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias TYPE(('e, 'f) IEEE.float))" using \mantissa f \ 0\ unfolding f_def bitlen_mantissa_Float using bitlen_denormal_mantissa[of x] by (auto simp: denormal_exponent_def) moreover have "2 - 2 ^ (LENGTH('e) - Suc 0) - int LENGTH('f) \ Float.exponent f" (is "?l \ _") proof - have "?l \ denormal_exponent TYPE(('e, 'f)float) + i" using that by (auto simp: is_denormal_def bias_def denormal_exponent_def of_nat_diff) also have "\ = Float.exponent f" unfolding i by auto finally show ?thesis . qed ultimately show ?thesis unfolding is_denormal_Float_def exponent_Float f_def[symmetric] by auto qed lemma is_finite_Float_of_finite: "is_finite_Float TYPE(('e, 'f)float) (of_finite x)" for x::"('e, 'f)float" by (auto simp: is_finite_Float_def of_finite_def is_normal_Float_normal is_denormal_Float_denormal) end diff --git a/thys/IEEE_Floating_Point/IEEE_Properties.thy b/thys/IEEE_Floating_Point/IEEE_Properties.thy --- a/thys/IEEE_Floating_Point/IEEE_Properties.thy +++ b/thys/IEEE_Floating_Point/IEEE_Properties.thy @@ -1,1023 +1,1022 @@ (* Author: Lei Yu, University of Cambridge Author: Fabian Hellauer Fabian Immler *) section \Proofs of Properties about Floating Point Arithmetic\ theory IEEE_Properties imports IEEE begin subsection \Theorems derived from definitions\ lemma valof_eq: "valof x = (if exponent x = 0 then (- 1) ^ sign x * (2 / 2 ^ bias TYPE(('a, 'b) float)) * (real (fraction x) / 2 ^ LENGTH('b)) else (- 1) ^ sign x * (2 ^ exponent x / 2 ^ bias TYPE(('a, 'b) float)) * (1 + real (fraction x) / 2 ^ LENGTH('b)))" for x::"('a, 'b) float" unfolding Let_def by transfer (auto simp: bias_def divide_simps unat_eq_0) lemma exponent_le[simp]: "exponent a \ unat (-1::'a word)" for a::"('a, _)float" by transfer (auto intro!: unat_mono simp: word_le_nat_alt[symmetric]) lemma max_word_le_exponent_iff[simp]: "unat (- 1::'a word) \ exponent a \ unat (- 1::'a word) = exponent a" for a::"('a, _)float" using le_antisym by fastforce lemma infinity_simps: "sign (plus_infinity::('e, 'f)float) = 0" "sign (minus_infinity::('e, 'f)float) = 1" "exponent (plus_infinity::('e, 'f)float) = emax TYPE(('e, 'f)float)" "exponent (minus_infinity::('e, 'f)float) = emax TYPE(('e, 'f)float)" "fraction (plus_infinity::('e, 'f)float) = 0" "fraction (minus_infinity::('e, 'f)float) = 0" subgoal by transfer auto subgoal by transfer auto subgoal by transfer (auto simp: emax_def) subgoal by transfer (auto simp: emax_def) subgoal by transfer auto subgoal by transfer auto done lemma zero_simps: "sign (0::('e, 'f)float) = 0" "sign (- 0::('e, 'f)float) = 1" "exponent (0::('e, 'f)float) = 0" "exponent (- 0::('e, 'f)float) = 0" "fraction (0::('e, 'f)float) = 0" "fraction (- 0::('e, 'f)float) = 0" subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto done lemma emax_eq: "emax x = 2^LENGTH('e) - 1" for x::"('e, 'f)float itself" by (simp add: emax_def unat_minus_one_word) lemma topfloat_simps: "sign (topfloat::('e, 'f)float) = 0" "exponent (topfloat::('e, 'f)float) = emax TYPE(('e, 'f)float) - 1" "fraction (topfloat::('e, 'f)float) = 2^fracwidth TYPE(('e, 'f)float) - 1" and bottomfloat_simps: "sign (bottomfloat::('e, 'f)float) = 1" "exponent (bottomfloat::('e, 'f)float) = emax TYPE(('e, 'f)float) - 1" "fraction (bottomfloat::('e, 'f)float) = 2^fracwidth TYPE(('e, 'f)float) - 1" subgoal by transfer simp subgoal by transfer (simp add: emax_eq take_bit_minus_small_eq nat_diff_distrib nat_power_eq) subgoal by transfer (simp add: unat_minus_one_word) subgoal by transfer simp subgoal by transfer (simp add: emax_eq take_bit_minus_small_eq unat_minus_one_word nat_diff_distrib nat_power_eq) subgoal by transfer (simp add: unat_minus_one_word) done lemmas float_defs = is_finite_def is_infinity_def is_zero_def is_nan_def is_normal_def is_denormal_def valof_eq less_eq_float_def less_float_def flt_def fgt_def fle_def fge_def feq_def fcompare_def infinity_simps zero_simps topfloat_simps bottomfloat_simps float_eq_def lemma float_cases: "is_nan a \ is_infinity a \ is_normal a \ is_denormal a \ is_zero a" by (auto simp: emax_def float_defs not_less) lemma float_cases_finite: "is_nan a \ is_infinity a \ is_finite a" by (simp add: float_cases is_finite_def) lemma float_zero1[simp]: "is_zero 0" unfolding float_defs by transfer auto lemma float_zero2[simp]: "is_zero (- x) \ is_zero x" unfolding float_defs by transfer auto lemma emax_pos[simp]: "0 < emax x" "emax x \ 0" by (auto simp: emax_def) text \The types of floating-point numbers are mutually distinct.\ lemma float_distinct: "\ (is_nan a \ is_infinity a)" "\ (is_nan a \ is_normal a)" "\ (is_nan a \ is_denormal a)" "\ (is_nan a \ is_zero a)" "\ (is_infinity a \ is_normal a)" "\ (is_infinity a \ is_denormal a)" "\ (is_infinity a \ is_zero a)" "\ (is_normal a \ is_denormal a)" "\ (is_denormal a \ is_zero a)" by (auto simp: float_defs) lemma denormal_imp_not_zero: "is_denormal f \ \is_zero f" by (simp add: is_denormal_def is_zero_def) lemma normal_imp_not_zero: "is_normal f \ \is_zero f" by (simp add: is_normal_def is_zero_def) lemma normal_imp_not_denormal: "is_normal f \ \is_denormal f" by (simp add: is_normal_def is_denormal_def) lemma denormal_zero[simp]: "\is_denormal 0" "\is_denormal minus_zero" using denormal_imp_not_zero float_zero1 float_zero2 by blast+ lemma normal_zero[simp]: "\is_normal 0" "\is_normal minus_zero" using normal_imp_not_zero float_zero1 float_zero2 by blast+ lemma float_distinct_finite: "\ (is_nan a \ is_finite a)" "\(is_infinity a \ is_finite a)" by (auto simp: float_defs) lemma finite_infinity: "is_finite a \ \ is_infinity a" by (auto simp: float_defs) lemma finite_nan: "is_finite a \ \ is_nan a" by (auto simp: float_defs) text \For every real number, the floating-point numbers closest to it always exist.\ lemma is_closest_exists: fixes v :: "('e, 'f)float \ real" and s :: "('e, 'f)float set" assumes finite: "finite s" and non_empty: "s \ {}" shows "\a. is_closest v s x a" using finite non_empty proof (induct s rule: finite_induct) case empty then show ?case by simp next case (insert z s) show ?case proof (cases "s = {}") case True then have "is_closest v (insert z s) x z" by (auto simp: is_closest_def) then show ?thesis by metis next case False then obtain a where a: "is_closest v s x a" using insert by metis then show ?thesis proof (cases "\v a - x\" "\v z - x\" rule: le_cases) case le then show ?thesis by (metis insert_iff a is_closest_def) next case ge have "\b. b \ s \ \v a - x\ \ \v b - x\" by (metis a is_closest_def) then have "\b. b \ insert z s \ \v z - x\ \ \v b - x\" by (metis eq_iff ge insert_iff order.trans) then show ?thesis using is_closest_def a by (metis insertI1) qed qed qed lemma closest_is_everything: fixes v :: "('e, 'f)float \ real" and s :: "('e, 'f)float set" assumes finite: "finite s" and non_empty: "s \ {}" shows "is_closest v s x (closest v p s x) \ ((\b. is_closest v s x b \ p b) \ p (closest v p s x))" unfolding closest_def by (rule someI_ex) (metis assms is_closest_exists [of s v x]) lemma closest_in_set: fixes v :: "('e, 'f)float \ real" assumes "finite s" and "s \ {}" shows "closest v p s x \ s" by (metis assms closest_is_everything is_closest_def) lemma closest_is_closest_finite: fixes v :: "('e, 'f)float \ real" assumes "finite s" and "s \ {}" shows "is_closest v s x (closest v p s x)" by (metis closest_is_everything assms) instance float::(len, len) finite proof qed (transfer, simp) lemma is_finite_nonempty: "{a. is_finite a} \ {}" proof - have "0 \ {a. is_finite a}" unfolding float_defs by transfer auto then show ?thesis by (metis empty_iff) qed lemma closest_is_closest: fixes v :: "('e, 'f)float \ real" assumes "s \ {}" shows "is_closest v s x (closest v p s x)" by (rule closest_is_closest_finite) (auto simp: assms) subsection \Properties about ordering and bounding\ text \Lifting of non-exceptional comparisons.\ lemma float_lt [simp]: assumes "is_finite a" "is_finite b" shows "a < b \ valof a < valof b" proof assume "valof a < valof b" moreover have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) ultimately have "fcompare a b = Lt" by (auto simp add: is_infinity_def is_nan_def valof_def fcompare_def) then show "a < b" by (auto simp: float_defs) next assume "a < b" then have lt: "fcompare a b = Lt" by (simp add: float_defs) have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then show "valof a < valof b" using lt assms by (simp add: fcompare_def is_nan_def is_infinity_def valof_def split: if_split_asm) qed lemma float_eq [simp]: assumes "is_finite a" "is_finite b" shows "a \ b \ valof a = valof b" proof assume *: "valof a = valof b" have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto with * have "fcompare a b = Eq" by (auto simp add: is_infinity_def is_nan_def valof_def fcompare_def) then show "a \ b" by (auto simp: float_defs) next assume "a \ b" then have eq: "fcompare a b = Eq" by (simp add: float_defs) have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto then show "valof a = valof b" using eq assms by (simp add: fcompare_def is_nan_def is_infinity_def valof_def split: if_split_asm) qed lemma float_le [simp]: assumes "is_finite a" "is_finite b" shows "a \ b \ valof a \ valof b" proof - have "a \ b \ a < b \ a \ b" by (auto simp add: float_defs) then show ?thesis by (auto simp add: assms) qed text \Reflexivity of equality for non-NaNs.\ lemma float_eq_refl [simp]: "a \ a \ \ is_nan a" by (auto simp: float_defs) text \Properties about Ordering.\ lemma float_lt_trans: "is_finite a \ is_finite b \ is_finite c \ a < b \ b < c \ a < c" by (auto simp: le_trans) lemma float_le_less_trans: "is_finite a \ is_finite b \ is_finite c \ a \ b \ b < c \ a < c" by (auto simp: le_trans) lemma float_le_trans: "is_finite a \ is_finite b \ is_finite c \ a \ b \ b \ c \ a \ c" by (auto simp: le_trans) lemma float_le_neg: "is_finite a \ is_finite b \ \ a < b \ b \ a" by auto text \Properties about bounding.\ lemma float_le_infinity [simp]: "\ is_nan a \ a \ plus_infinity" unfolding float_defs by auto lemma zero_le_topfloat[simp]: "0 \ topfloat" "- 0 \ topfloat" by (auto simp: float_defs field_simps power_gt1_lemma of_nat_diff) lemma LENGTH_contr: "Suc 0 < LENGTH('e) \ 2 ^ LENGTH('e::len) \ Suc (Suc 0) \ False" by (metis le_antisym len_gt_0 n_less_equal_power_2 not_less_eq numeral_2_eq_2 one_le_numeral self_le_power) lemma valof_topfloat: "valof (topfloat::('e, 'f)float) = largest TYPE(('e, 'f)float)" if "LENGTH('e) > 1" using that LENGTH_contr by (auto simp add: emax_eq largest_def divide_simps float_defs of_nat_diff) lemma float_frac_le: "fraction a \ 2^LENGTH('f) - 1" for a::"('e, 'f)float" unfolding float_defs using less_Suc_eq_le by transfer fastforce lemma float_exp_le: "is_finite a \ exponent a \ emax TYPE(('e, 'f)float) - 1" for a::"('e, 'f)float" unfolding float_defs by auto lemma float_sign_le: "(-1::real)^(sign a) = 1 \ (-1::real)^(sign a) = -1" by (metis neg_one_even_power neg_one_odd_power) lemma exp_less: "a \ b \ (2::real)^a \ 2^b" for a b :: nat by auto lemma div_less: "a \ b \ c > 0 \ a/c \ b/c" for a b c :: "'a::linordered_field" by (metis divide_le_cancel less_asym) lemma finite_topfloat: "is_finite topfloat" unfolding float_defs by auto lemmas float_leI = float_le[THEN iffD2] lemma factor_minus: "x * a - x = x * (a - 1)" for x a::"'a::comm_semiring_1_cancel" by (simp add: algebra_simps) lemma real_le_power_numeral_diff: "real a \ numeral b ^ n - 1 \ a \ numeral b ^ n - 1" by (metis (mono_tags, lifting) of_nat_1 of_nat_diff of_nat_le_iff of_nat_numeral one_le_numeral one_le_power semiring_1_class.of_nat_power) definition denormal_exponent::"('e, 'f)float itself \ int" where "denormal_exponent x = 1 - (int (LENGTH('f)) + int (bias TYPE(('e, 'f)float)))" definition normal_exponent::"('e, 'f)float \ int" where "normal_exponent x = int (exponent x) - int (bias TYPE(('e, 'f)float)) - int (LENGTH('f))" definition denormal_mantissa::"('e, 'f)float \ int" where "denormal_mantissa x = (-1::int)^sign x * int (fraction x)" definition normal_mantissa::"('e, 'f)float \ int" where "normal_mantissa x = (-1::int)^sign x * (2^LENGTH('f) + int (fraction x))" lemma unat_one_word_le: "unat a \ Suc 0" for a::"1 word" using unat_lt2p[of a] by auto lemma one_word_le: "a \ 1" for a::"1 word" by (auto simp: word_le_nat_alt unat_one_word_le) lemma sign_cases[case_names pos neg]: obtains "sign x = 0" | "sign x = 1" proof cases assume "sign x = 0" then show ?thesis .. next assume "sign x \ 0" have "sign x \ 1" by transfer (auto simp: unat_one_word_le) then have "sign x = 1" using \sign x \ 0\ by auto then show ?thesis .. qed lemma is_infinity_cases: assumes "is_infinity x" obtains "x = plus_infinity" | "x = minus_infinity" proof (cases rule: sign_cases[of x]) assume "sign x = 0" then have "x = plus_infinity" using assms unfolding float_defs - by transfer (auto simp: emax_def unat_eq_0) + by transfer (auto simp: emax_def unat_eq_0 simp flip: word_eq_iff_unsigned) then show ?thesis .. next assume "sign x = 1" then have "x = minus_infinity" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_of_nat) then show ?thesis .. qed lemma is_zero_cases: assumes "is_zero x" obtains "x = 0" | "x = - 0" proof (cases rule: sign_cases[of x]) assume "sign x = 0" then have "x = 0" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_0) then show ?thesis .. next assume "sign x = 1" then have "x = minus_zero" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_of_nat) then show ?thesis .. qed lemma minus_minus_float[simp]: "- (-f) = f" for f::"('e, 'f)float" by transfer auto lemma sign_minus_float: "sign (-f) = (1 - sign f)" for f::"('e, 'f)float" by transfer (auto simp: unat_eq_1 one_word_le unat_sub) lemma exponent_uminus[simp]: "exponent (- f) = exponent f" by transfer auto lemma fraction_uminus[simp]: "fraction (- f) = fraction f" by transfer auto lemma is_normal_minus_float[simp]: "is_normal (-f) = is_normal f" for f::"('e, 'f)float" by (auto simp: is_normal_def) lemma is_denormal_minus_float[simp]: "is_denormal (-f) = is_denormal f" for f::"('e, 'f)float" by (auto simp: is_denormal_def) lemma bitlen_normal_mantissa: "bitlen (abs (normal_mantissa x)) = Suc LENGTH('f)" for x::"('e, 'f)float" proof - have "fraction x < 2 ^ LENGTH('f)" using float_frac_le[of x] by (metis One_nat_def Suc_pred le_imp_less_Suc pos2 zero_less_power) moreover have "- int (fraction x) \ 2 ^ LENGTH('f)" using negative_zle_0 order_trans zero_le_numeral zero_le_power by blast ultimately show ?thesis by (cases x rule: sign_cases) (auto simp: bitlen_le_iff_power bitlen_ge_iff_power nat_add_distrib normal_mantissa_def intro!: antisym) qed lemma less_int_natI: "x < y" if "0 \ x" "nat x < nat y" using that by arith lemma normal_exponent_bounds_int: "2 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) \ normal_exponent x" "normal_exponent x \ 2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1" if "is_normal x" for x::"('e, 'f)float" using that unfolding normal_exponent_def is_normal_def emax_eq bias_def by (auto simp del: zless_nat_conj intro!: less_int_natI simp: of_nat_diff nat_add_distrib nat_mult_distrib nat_power_eq power_Suc[symmetric]) lemmas of_int_leI = of_int_le_iff[THEN iffD2] lemma normal_exponent_bounds_real: "2 - 2 ^ (LENGTH('e) - 1) - real LENGTH('f) \ normal_exponent x" "normal_exponent x \ 2 ^ (LENGTH('e) - 1) - real LENGTH('f) - 1" if "is_normal x" for x::"('e, 'f)float" subgoal by (rule order_trans[OF _ of_int_leI[OF normal_exponent_bounds_int(1)[OF that]]]) auto subgoal by (rule order_trans[OF of_int_leI[OF normal_exponent_bounds_int(2)[OF that]]]) auto done lemma float_eqI: "x = y" if "sign x = sign y" "fraction x = fraction y" "exponent x = exponent y" - using that - by transfer auto + using that by transfer (auto simp add: word_unat_eq_iff) lemma float_induct[induct type:float, case_names normal denormal neg zero infinity nan]: fixes a::"('e, 'f)float" assumes normal: "\x. is_normal x \ valof x = normal_mantissa x * 2 powr normal_exponent x \ P x" assumes denormal: "\x. is_denormal x \ valof x = denormal_mantissa x * 2 powr denormal_exponent TYPE(('e, 'f)float) \ P x" assumes zero: "P 0" "P minus_zero" assumes infty: "P plus_infinity" "P minus_infinity" assumes nan: "\x. is_nan x \ P x" shows "P a" proof - from float_cases[of a] consider "is_nan a" | "is_infinity a" | "is_normal a" | "is_denormal a" | "is_zero a" by blast then show ?thesis proof cases case 1 then show ?thesis by (rule nan) next case 2 then consider "a = plus_infinity" | "a = minus_infinity" by (rule is_infinity_cases) then show ?thesis by cases (auto intro: infty) next case hyps: 3 from hyps have "valof a = normal_mantissa a * 2 powr normal_exponent a" by (cases a rule: sign_cases) (auto simp: valof_eq normal_mantissa_def normal_exponent_def is_normal_def powr_realpow[symmetric] powr_diff powr_add divide_simps algebra_simps) from hyps this show ?thesis by (rule normal) next case hyps: 4 from hyps have "valof a = denormal_mantissa a * 2 powr denormal_exponent TYPE(('e, 'f)float)" by (cases a rule: sign_cases) (auto simp: valof_eq denormal_mantissa_def denormal_exponent_def is_denormal_def powr_realpow[symmetric] powr_diff powr_add divide_simps algebra_simps) from hyps this show ?thesis by (rule denormal) next case 5 then consider "a = 0" | "a = minus_zero" by (rule is_zero_cases) then show ?thesis by cases (auto intro: zero) qed qed lemma infinite_infinity[simp]: "\ is_finite plus_infinity" "\ is_finite minus_infinity" by (auto simp: is_finite_def is_normal_def infinity_simps is_denormal_def is_zero_def) lemma nan_not_finite[simp]: "is_nan x \ \ is_finite x" using float_distinct_finite(1) by blast lemma valof_nonneg: "valof x \ 0" if "sign x = 0" for x::"('e, 'f)float" by (auto simp: valof_eq that) lemma valof_nonpos: "valof x \ 0" if "sign x = 1" for x::"('e, 'f)float" using that by (auto simp: valof_eq is_finite_def) lemma real_le_intI: "x \ y" if "floor x \ floor y" "x \ \" for x y::real using that(2,1) by (induction rule: Ints_induct) (auto elim!: Ints_induct simp: le_floor_iff) lemma real_of_int_le_2_powr_bitlenI: "real_of_int x \ 2 powr n - 1" if "bitlen (abs x) \ m" "m \ n" proof - have "real_of_int x \ abs (real_of_int x)" by simp also have "\ < 2 powr (bitlen (abs x))" by (rule abs_real_le_2_powr_bitlen) finally have "real_of_int x \ 2 powr (bitlen (abs x)) - 1" by (auto simp: powr_real_of_int bitlen_nonneg intro!: real_le_intI) also have "\ \ 2 powr m - 1" by (simp add: that) also have "\ \ 2 powr n - 1" by (simp add: that) finally show ?thesis . qed lemma largest_eq: "largest TYPE(('e, 'f)float) = (2 ^ (LENGTH('f) + 1) - 1) * 2 powr real_of_int (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" proof - have "2 ^ LENGTH('e) - 1 - 1 = (2::nat) ^ LENGTH('e) - 2" by arith then have "largest TYPE(('e, 'f)float) = (2 ^ (LENGTH('f) + 1) - 1) * 2 powr (real (2 ^ LENGTH('e) - 2) + 1 - real (2 ^ (LENGTH('e) - 1)) - LENGTH('f))" unfolding largest_def emax_eq bias_def by (auto simp: largest_def powr_realpow[symmetric] powr_minus divide_simps algebra_simps powr_diff powr_add of_nat_diff) also have "2 ^ LENGTH('e) \ (2::nat)" by (simp add: self_le_power) then have "(real (2 ^ LENGTH('e) - 2) + 1 - real (2 ^ (LENGTH('e) - 1)) - LENGTH('f)) = (real (2 ^ LENGTH('e)) - 2 ^ (LENGTH('e) - 1) - LENGTH('f)) - 1" by (auto simp add: of_nat_diff) also have "real (2 ^ LENGTH('e)) = 2 ^ LENGTH('e)" by auto also have "(2 ^ LENGTH('e) - 2 ^ (LENGTH('e) - 1) - real LENGTH('f) - 1) = real_of_int ((2 ^ (LENGTH('e) - 1) - int (LENGTH('f)) - 1))" by (simp, subst power_Suc[symmetric], simp) finally show ?thesis by simp qed lemma bitlen_denormal_mantissa: "bitlen (abs (denormal_mantissa x)) \ LENGTH('f)" for x::"('e, 'f)float" proof - have "fraction x < 2 ^ LENGTH('f)" using float_frac_le[of x] by (metis One_nat_def Suc_pred le_imp_less_Suc pos2 zero_less_power) moreover have "- int (fraction x) \ 2 ^ LENGTH('f)" using negative_zle_0 order_trans zero_le_numeral zero_le_power by blast ultimately show ?thesis by (cases x rule: sign_cases) (auto simp: bitlen_le_iff_power denormal_mantissa_def intro!: antisym) qed lemma float_le_topfloat: fixes a::"('e, 'f)float" assumes "is_finite a" "LENGTH('e) > 1" shows "a \ topfloat" using assms(1) proof (induction a rule: float_induct) case (normal x) note normal(2) also have "real_of_int (normal_mantissa x) * 2 powr real_of_int (normal_exponent x) \ (2 powr (LENGTH('f) + 1) - 1) * 2 powr (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" using normal_exponent_bounds_real(2)[OF \is_normal x\] by (auto intro!: mult_mono real_of_int_le_2_powr_bitlenI simp: bitlen_normal_mantissa powr_realpow[symmetric] ge_one_powr_ge_zero) also have "\ = largest TYPE(('e, 'f) IEEE.float)" unfolding largest_eq by (auto simp: powr_realpow powr_add) also have "\ = valof (topfloat::('e, 'f) float)" using assms by (simp add: valof_topfloat) finally show ?case by (intro float_leI normal finite_topfloat) next case (denormal x) note denormal(2) also have "3 \ 2 powr (1 + real (LENGTH('e) - Suc 0))" proof - have "3 \ 2 powr (2::real)" by simp also have "\ \ 2 powr (1 + real (LENGTH('e) - Suc 0))" using assms by (subst powr_le_cancel_iff) auto finally show ?thesis . qed then have "real_of_int (denormal_mantissa x) * 2 powr real_of_int (denormal_exponent TYPE(('e, 'f)float)) \ (2 powr (LENGTH('f) + 1) - 1) * 2 powr (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" using bitlen_denormal_mantissa[of x] by (auto intro!: mult_mono real_of_int_le_2_powr_bitlenI simp: bitlen_normal_mantissa powr_realpow[symmetric] ge_one_powr_ge_zero denormal_exponent_def bias_def powr_mult_base of_nat_diff) also have "\ \ largest TYPE(('e, 'f) IEEE.float)" unfolding largest_eq by (rule mult_mono) (auto simp: powr_realpow powr_add power_Suc[symmetric] simp del: power_Suc) also have "\ = valof (topfloat::('e, 'f) float)" using assms by (simp add: valof_topfloat) finally show ?case by (intro float_leI denormal finite_topfloat) qed auto lemma float_val_le_largest: "valof a \ largest TYPE(('e, 'f)float)" if "is_finite a" "LENGTH('e) > 1" for a::"('e, 'f)float" by (metis that finite_topfloat float_le float_le_topfloat valof_topfloat) lemma float_val_lt_threshold: "valof a < threshold TYPE(('e, 'f)float)" if "is_finite a" "LENGTH('e) > 1" for a::"('e, 'f)float" proof - have "valof a \ largest TYPE(('e, 'f)float)" by (rule float_val_le_largest [OF that]) also have "\ < threshold TYPE(('e, 'f)float)" by (auto simp: largest_def threshold_def divide_simps) finally show ?thesis . qed subsection \Algebraic properties about basic arithmetic\ text \Commutativity of addition.\ lemma assumes "is_finite a" "is_finite b" shows float_plus_comm_eq: "a + b = b + a" and float_plus_comm: "is_finite (a + b) \ (a + b) \ (b + a)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then show "a + b = b + a" by (simp add: float_defs fadd_def plus_float_def add.commute) then show "is_finite (a + b) \ (a + b) \ (b + a)" by (metis float_eq) qed text \The floating-point number a falls into the same category as the negation of \a\.\ lemma is_zero_uminus[simp]: "is_zero (- a) \ is_zero a" by (simp add: is_zero_def) lemma is_infinity_uminus [simp]: "is_infinity (- a) = is_infinity a" by (simp add: is_infinity_def) lemma is_finite_uminus[simp]: "is_finite (- a) \ is_finite a" by (simp add: is_finite_def) lemma is_nan_uminus[simp]: "is_nan (- a) \ is_nan a" by (simp add: is_nan_def) text \The sign of a and the sign of a's negation is different.\ lemma float_neg_sign: "(sign a) \ (sign (- a))" by (cases a rule: sign_cases) (auto simp: sign_minus_float) lemma float_neg_sign1: "sign a = sign (- b) \ sign a \ sign b" by (metis float_neg_sign sign_cases) lemma valof_uminus: assumes "is_finite a" shows "valof (- a) = - valof a" (is "?L = ?R") by (cases a rule: sign_cases) (auto simp: valof_eq sign_minus_float) text \Showing \a + (- b) = a - b\.\ lemma float_neg_add: "is_finite a \ is_finite b \ is_finite (a - b) \ valof a + valof (- b) = valof a - valof b" by (simp add: valof_uminus) lemma float_plus_minus: assumes "is_finite a" "is_finite b" "is_finite (a - b)" shows "(a + - b) \ (a - b)" proof - have nab: "\ is_nan a" "\ is_nan (- b)" "\ is_infinity a" "\ is_infinity (- b)" using assms by (auto simp: finite_nan finite_infinity) have "a - b = (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then (sign a) else 0) (round To_nearest (valof a - valof b)))" using nab by (auto simp: minus_float_def fsub_def) also have "\ = (zerosign (if is_zero a \ is_zero (- b) \ sign a = sign (- b) then sign a else 0) (round To_nearest (valof a + valof (- b))))" using assms by (simp add: float_neg_sign1 float_neg_add) also have "\ = a + - b" using nab by (auto simp: float_defs fadd_def plus_float_def) finally show ?thesis using assms by (metis float_eq) qed lemma finite_bottomfloat: "is_finite bottomfloat" by (simp add: finite_topfloat) lemma bottomfloat_eq_m_largest: "valof (bottomfloat::('e, 'f)float) = - largest TYPE(('e, 'f)float)" if "LENGTH('e) > 1" using that by (auto simp: valof_topfloat valof_uminus finite_topfloat) lemma float_val_ge_bottomfloat: "valof a \ valof (bottomfloat::('e, 'f)float)" if "LENGTH('e) > 1" "is_finite a" for a::"('e,'f)float" proof - have "- a \ topfloat" using that by (auto intro: float_le_topfloat) then show ?thesis using that by (auto simp: valof_uminus finite_topfloat) qed lemma float_ge_bottomfloat: "is_finite a \ a \ bottomfloat" if "LENGTH('e) > 1" "is_finite a" for a::"('e,'f)float" by (metis finite_bottomfloat float_le float_val_ge_bottomfloat that) lemma float_val_ge_largest: fixes a::"('e,'f)float" assumes "LENGTH('e) > 1" "is_finite a" shows "valof a \ - largest TYPE(('e,'f)float)" proof - have "- largest TYPE(('e,'f)float) = valof (bottomfloat::('e,'f)float)" using assms by (simp add: bottomfloat_eq_m_largest) also have "\ \ valof a" using assms by (simp add: float_val_ge_bottomfloat) finally show ?thesis . qed lemma float_val_gt_threshold: fixes a::"('e,'f)float" assumes "LENGTH('e) > 1" "is_finite a" shows "valof a > - threshold TYPE(('e,'f)float)" proof - have largest: "valof a \ -largest TYPE(('e,'f)float)" using assms by (metis float_val_ge_largest) then have "-largest TYPE(('e,'f)float) > - threshold TYPE(('e,'f)float)" by (auto simp: bias_def threshold_def largest_def divide_simps) then show ?thesis by (metis largest less_le_trans) qed text \Showing \abs (- a) = abs a\.\ lemma float_abs [simp]: "\ is_nan a \ abs (- a) = abs a" by (metis IEEE.abs_float_def float_neg_sign1 minus_minus_float zero_simps(1)) lemma neg_zerosign: "- (zerosign s a) = zerosign (1 - s) (- a)" by (auto simp: zerosign_def) subsection \Properties about Rounding Errors\ definition error :: "('e, 'f)float itself \ real \ real" where "error _ x = valof (round To_nearest x::('e, 'f)float) - x" lemma bound_at_worst_lemma: fixes a::"('e, 'f)float" assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" assumes finite: "is_finite a" shows "\valof (round To_nearest x::('e, 'f)float) - x\ \ \valof a - x\" proof - have *: "(round To_nearest x::('e,'f)float) = closest valof (\a. even (fraction a)) {a. is_finite a} x" using threshold finite by auto have "is_closest (valof) {a. is_finite a} x (round To_nearest x::('e,'f)float)" using is_finite_nonempty unfolding * by (intro closest_is_closest) auto then show ?thesis using finite is_closest_def by (metis mem_Collect_eq) qed lemma error_at_worst_lemma: fixes a::"('e, 'f)float" assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" and "is_finite a" shows "\error TYPE(('e, 'f)float) x\ \ \valof a - x\ " unfolding error_def by (rule bound_at_worst_lemma; fact) lemma error_is_zero [simp]: fixes a::"('e, 'f)float" assumes "is_finite a" "1 < LENGTH('e)" shows "error TYPE(('e, 'f)float) (valof a) = 0" proof - have "\valof a\ < threshold TYPE(('e, 'f)float)" by (metis abs_less_iff minus_less_iff float_val_gt_threshold float_val_lt_threshold assms) then show ?thesis by (metis abs_le_zero_iff abs_zero diff_self error_at_worst_lemma assms(1)) qed lemma is_finite_zerosign[simp]: "is_finite (zerosign s a) \ is_finite a" by (auto simp: zerosign_def is_finite_def) lemma is_finite_closest: "is_finite (closest (v::_\real) p {a. is_finite a} x)" using closest_is_closest[OF is_finite_nonempty, of v x p] by (auto simp: is_closest_def) lemma defloat_float_zerosign_round_finite: assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" shows "is_finite (zerosign s (round To_nearest x::('e, 'f)float))" proof - have "(round To_nearest x::('e, 'f)float) = (closest valof (\a. even (fraction a)) {a. is_finite a} x)" using threshold by (metis (full_types) abs_less_iff leD le_minus_iff round.simps(1)) then have "is_finite (round To_nearest x::('e, 'f)float)" by (metis is_finite_closest) then show ?thesis using is_finite_zerosign by auto qed lemma valof_zero[simp]: "valof 0 = 0" "valof minus_zero = 0" by (auto simp add: zerosign_def valof_eq zero_simps) lemma signzero_zero: "is_zero a \ valof (zerosign s a) = 0" by (auto simp add: zerosign_def) lemma val_zero: "is_zero a \ valof a = 0" by (cases a rule: is_zero_cases) auto lemma float_add: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a + valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_add: "is_finite (a + b)" and error_float_add: "valof (a + b) = valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto then have ab: "(a + b) = (zerosign (if is_zero a \ is_zero b \ sign a = sign b then (sign a) else 0) (round To_nearest (valof a + valof b)))" using assms by (auto simp add: float_defs fadd_def plus_float_def) then show "is_finite ((a + b))" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a + b) = valof (zerosign (if is_zero a \ is_zero b \ sign a = sign b then (sign a) else 0) (round To_nearest (valof a + valof b)::('e, 'f)float))" by (auto simp: ab is_infinity_def is_nan_def valof_def) show "valof (a + b) = valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b)" proof (cases "is_zero (round To_nearest (valof a + valof b)::('e, 'f)float)") case True have "valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b) = valof (round To_nearest (valof a + valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_sub: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a - valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_sub: "is_finite (a - b)" and error_float_sub: "valof (a - b) = valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then have ab: "a - b = (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else 0) (round To_nearest (valof a - valof b)))" using assms by (auto simp add: float_defs fsub_def minus_float_def) then show "is_finite (a - b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a - b) = valof (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else 0) (round To_nearest (valof a - valof b)::('e, 'f)float))" by (auto simp: ab is_infinity_def is_nan_def valof_def) show "valof (a - b) = valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b)" proof (cases "is_zero (round To_nearest (valof a - valof b)::('e, 'f)float)") case True have "valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b) = valof (round To_nearest (valof a - valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_mul: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a * valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_mul: "is_finite (a * b)" and error_float_mul: "valof (a * b) = valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b)" proof - have non: "\ is_nan a" "\ is_nan b" "\ is_infinity a" "\ is_infinity b" using assms float_distinct_finite by auto then have ab: "a * b = (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a * valof b)::('e, 'f)float))" using assms by (auto simp: float_defs fmul_def times_float_def) then show "is_finite (a * b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a * b) = valof (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a * valof b)::('e, 'f)float))" by (auto simp: ab float_defs of_bool_def) show "valof (a * b) = valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b)" proof (cases "is_zero (round To_nearest (valof a * valof b)::('e, 'f)float)") case True have "valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b) = valof (round To_nearest (valof a * valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_div: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and not_zero: "\ is_zero b" and threshold: "\valof a / valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_div: "is_finite (a / b)" and error_float_div: "valof (a / b) = valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b)" proof - have ab: "a / b = (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a / valof b)))" using assms by (simp add: divide_float_def fdiv_def finite_infinity finite_nan not_zero float_defs [symmetric]) then show "is_finite (a / b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a / b) = valof (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a / valof b))::('e, 'f)float)" by (auto simp: ab float_defs of_bool_def) show "valof (a / b) = valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b)" proof (cases "is_zero (round To_nearest (valof a / valof b)::('e, 'f)float)") case True have "valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b) = valof (round To_nearest (valof a / valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma valof_one[simp]: "valof (1::('e, 'f)float) = (if LENGTH('e) \ 1 then 0 else 1)" by transfer (auto simp: bias_def unat_sub word_1_le_power p2_eq_1) end diff --git a/thys/Iptables_Semantics/Common/Word_Upto.thy b/thys/Iptables_Semantics/Common/Word_Upto.thy --- a/thys/Iptables_Semantics/Common/Word_Upto.thy +++ b/thys/Iptables_Semantics/Common/Word_Upto.thy @@ -1,206 +1,206 @@ section\Word Upto\ theory Word_Upto imports Main IP_Addresses.Hs_Compat Word_Lib.Word_Lemmas begin text\Enumerate a range of machine words.\ text\enumerate from the back (inefficient)\ function word_upto :: "'a word \ 'a word \ ('a::len) word list" where "word_upto a b = (if a = b then [a] else word_upto a (b - 1) @ [b])" by pat_completeness auto (*by the way: does not terminate practically if b < a; will terminate after it reaches the word wrap-around!*) termination word_upto apply(relation "measure (unat \ uncurry (-) \ prod.swap)") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply(simp add: diff_right_commute; fail) apply(rule measure_unat) apply auto done declare word_upto.simps[simp del] text\enumerate from the front (more inefficient)\ function word_upto' :: "'a word \ 'a word \ ('a::len) word list" where "word_upto' a b = (if a = b then [a] else a # word_upto' (a + 1) b)" by pat_completeness auto termination word_upto' apply(relation "measure (\ (a, b). unat (b - a))") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply (simp add: diff_diff_add; fail) apply(rule measure_unat) apply auto done declare word_upto'.simps[simp del] lemma word_upto_cons_front[code]: "word_upto a b = word_upto' a b" proof(induction a b rule:word_upto'.induct) case (1 a b) have hlp1: "a \ b \ a # word_upto (a + 1) b = word_upto a b" apply(induction a b rule:word_upto.induct) apply simp apply(subst(1) word_upto.simps) apply(simp) apply safe apply(subst(1) word_upto.simps) apply (simp) apply(subst(1) word_upto.simps) apply (simp; fail) apply(case_tac "a \ b - 1") apply(simp) apply (metis Cons_eq_appendI word_upto.simps) apply(simp) done from 1[symmetric] show ?case apply(cases "a = b") subgoal apply(subst word_upto.simps) apply(subst word_upto'.simps) by(simp) apply(subst word_upto'.simps) by(simp add: hlp1) qed (* Most of the lemmas I show about word_upto hold without a \ b, but I don't need that right now and it's giving me a headache *) lemma word_upto_set_eq: "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" proof show "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst(asm) word_upto.simps) apply(simp; fail) apply(subst(asm) word_upto.simps) apply(simp) apply(erule disjE) apply(simp; fail) proof(goal_cases) case (1 a b) from 1(2-3) have "b \ 0" by force from 1(2,3) have "a \ b - 1" by (simp add: word_le_minus_one_leq) from 1(1)[OF this 1(4)] show ?case by (metis dual_order.trans 1(2,3) less_imp_le measure_unat word_le_0_iff word_le_nat_alt) qed next show "a \ x \ x \ b \ x \ set (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(simp) apply(case_tac "x = b") apply(simp;fail) proof(goal_cases) case (1 a b) from 1(2-4) have "b \ 0" by force from 1(2,4) have "x \ b - 1" using le_step_down_word by auto from 1(1) this show ?case by simp qed qed lemma word_upto_distinct_hlp: "a \ b \ a \ b \ b \ set (word_upto a (b - 1))" apply(rule ccontr, unfold not_not) apply(subgoal_tac "a \ b - 1") apply(drule iffD1[OF word_upto_set_eq[of a "b -1" b]]) apply(simp add: word_upto.simps) apply(subgoal_tac "b \ 0") apply(meson leD measure_unat word_le_nat_alt) apply(blast intro: iffD1[OF word_le_0_iff]) using le_step_down_word apply blast done lemma distinct_word_upto: "a \ b \ distinct (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(case_tac "a \ b - 1") apply(simp) apply(rule word_upto_distinct_hlp; simp) apply(simp) apply(rule ccontr) apply (simp add: not_le antisym word_minus_one_le_leq) done lemma word_upto_eq_upto: "s \ e \ e \ unat (max_word :: 'l word) \ word_upto ((of_nat :: nat \ ('l :: len) word) s) (of_nat e) = map of_nat (upt s (Suc e))" proof(induction e) let ?mwon = "of_nat :: nat \ 'l word" let ?mmw = "max_word :: 'l word" case (Suc e) show ?case proof(cases "?mwon s = ?mwon (Suc e)") case True have "s = Suc e" using le_unat_uoi Suc.prems True by metis with True show ?thesis by(subst word_upto.simps) (simp) next case False hence le: "s \ e" using le_SucE Suc.prems by blast have lm: "e \ unat ?mmw" using Suc.prems by simp have sucm: "(of_nat :: nat \ ('l :: len) word) (Suc e) - 1 = of_nat e" using Suc.prems(2) by simp note mIH = Suc.IH[OF le lm] show ?thesis by(subst word_upto.simps) (simp add: False[simplified] Suc.prems mIH sucm) qed qed(simp add: word_upto.simps) lemma word_upto_alt: "(a :: ('l :: len) word) \ b \ word_upto a b = map of_nat (upt (unat a) (Suc (unat b)))" proof - let ?mmw = "- 1 :: 'l word" assume le: "a \ b" hence nle: "unat a \ unat b" by(unat_arith) have lem: "unat b \ unat ?mmw" by (simp add: word_unat_less_le) - note word_upto_eq_upto[OF nle lem, unfolded word_unat.Rep_inverse] - thus "word_upto a b = map of_nat [unat a.. b then map of_nat (upt (unat a) (Suc (unat b))) else word_upto a b)" using word_upto_alt by metis lemma sorted_word_upto: fixes a b :: "('l :: len) word" assumes "a \ b" shows "sorted (word_upto a b)" proof - define m and n where \m = unat a\ and \n = Suc (unat b)\ moreover have \sorted (map of_nat [m.. apply (simp add: sorted_map) apply (rule sorted_wrt_mono_rel [of _ \(\)\]) apply simp_all apply (simp add: le_unat_uoi less_Suc_eq_le n_def word_of_nat_le) apply transfer apply simp apply (subst take_bit_int_eq_self) apply (simp_all add: le_less_trans) apply (metis le_unat_uoi of_int_of_nat_eq of_nat_mono uint_word_of_int_eq unat_eq_nat_uint unsigned_of_int) done ultimately have \sorted (map of_nat [unat a.. by simp with assms show ?thesis by (simp only: word_upto_alt) qed end diff --git a/thys/Word_Lib/Aligned.thy b/thys/Word_Lib/Aligned.thy --- a/thys/Word_Lib/Aligned.thy +++ b/thys/Word_Lib/Aligned.thy @@ -1,1282 +1,1328 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports "HOL-Library.Word" More_Word Word_EqI - Typedef_Morphisms begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ is \\k n. 2 ^ n dvd take_bit LENGTH('a) k\ by simp lemma is_aligned_iff_udvd: \is_aligned w n \ 2 ^ n udvd w\ by transfer (simp flip: take_bit_eq_0_iff add: min_def) lemma is_aligned_iff_take_bit_eq_0: \is_aligned w n \ take_bit n w = 0\ by (simp add: is_aligned_iff_udvd take_bit_eq_0_iff exp_dvd_iff_exp_udvd) lemma is_aligned_iff_dvd_int: \is_aligned ptr n \ 2 ^ n dvd uint ptr\ by transfer simp lemma is_aligned_iff_dvd_nat: \is_aligned ptr n \ 2 ^ n dvd unat ptr\ proof - have \unat ptr = nat \uint ptr\\ by transfer simp then have \2 ^ n dvd unat ptr \ 2 ^ n dvd uint ptr\ by (simp only: dvd_nat_abs_iff) simp then show ?thesis by (simp add: is_aligned_iff_dvd_int) qed lemma is_aligned_0 [simp]: \is_aligned 0 n\ by transfer simp lemma is_aligned_at_0 [simp]: \is_aligned w 0\ by transfer simp lemma is_aligned_beyond_length: \is_aligned w n \ w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that apply (simp add: is_aligned_iff_udvd) apply transfer apply auto done lemma is_alignedI [intro?]: \is_aligned x n\ if \x = 2 ^ n * k\ for x :: \'a::len word\ proof (unfold is_aligned_iff_udvd) from that show \2 ^ n udvd x\ using dvd_triv_left exp_dvd_iff_exp_udvd by blast qed lemma is_alignedE: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ proof (cases \n < LENGTH('a)\) case False with assms have \w = 0\ by (simp add: is_aligned_beyond_length) with that [of 0] show thesis by simp next case True moreover define m where \m = LENGTH('a) - n\ ultimately have l: \LENGTH('a) = n + m\ and \m \ 0\ by simp_all from \n < LENGTH('a)\ have *: \unat (2 ^ n :: 'a word) = 2 ^ n\ by transfer simp from assms have \2 ^ n udvd w\ by (simp add: is_aligned_iff_udvd) then obtain v :: \'a word\ where \unat w = unat (2 ^ n :: 'a word) * unat v\ .. moreover define q where \q = unat v\ ultimately have unat_w: \unat w = 2 ^ n * q\ by (simp add: *) then have \word_of_nat (unat w) = (word_of_nat (2 ^ n * q) :: 'a word)\ by simp then have w: \w = 2 ^ n * word_of_nat q\ by simp moreover have \q < 2 ^ (LENGTH('a) - n)\ proof (rule ccontr) assume \\ q < 2 ^ (LENGTH('a) - n)\ then have \2 ^ (LENGTH('a) - n) \ q\ by simp then have \2 ^ LENGTH('a) \ 2 ^ n * q\ by (simp add: l power_add) with unat_w [symmetric] show False by (metis le_antisym nat_less_le unsigned_less) qed ultimately show thesis using that by blast qed lemma is_alignedE' [elim?]: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ proof - from assms obtain q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ by (rule is_alignedE) then have \w = push_bit n (word_of_nat q)\ by (simp add: push_bit_eq_mult) with that show thesis using \q < 2 ^ (LENGTH('a) - n)\ . qed lemma is_aligned_mask: \is_aligned w n \ w AND mask n = 0\ by (simp add: is_aligned_iff_take_bit_eq_0 take_bit_eq_mask) lemma is_aligned_imp_not_bit: \\ bit w m\ if \is_aligned w n\ and \m < n\ for w :: \'a::len word\ proof - from \is_aligned w n\ obtain q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ .. moreover have \\ bit (push_bit n (word_of_nat q :: 'a word)) m\ using \m < n\ by (simp add: bit_simps) ultimately show ?thesis by simp qed lemma is_aligned_weaken: "\ is_aligned w x; x \ y \ \ is_aligned w y" unfolding is_aligned_iff_dvd_nat by (erule dvd_trans [rotated]) (simp add: le_imp_power_dvd) lemma is_alignedE_pre: fixes w::"'a::len word" assumes aligned: "is_aligned w n" shows rl: "\q. w = 2 ^ n * (of_nat q) \ q < 2 ^ (LENGTH('a) - n)" using aligned is_alignedE by blast lemma aligned_add_aligned: fixes x::"'a::len word" assumes aligned1: "is_aligned x n" and aligned2: "is_aligned y m" and lt: "m \ n" shows "is_aligned (x + y) m" proof cases assume nlt: "n < LENGTH('a)" show ?thesis unfolding is_aligned_iff_dvd_nat dvd_def proof - from aligned2 obtain q2 where yv: "y = 2 ^ m * of_nat q2" and q2v: "q2 < 2 ^ (LENGTH('a) - m)" by (auto elim: is_alignedE) from lt obtain k where kv: "m + k = n" by (auto simp: le_iff_add) with aligned1 obtain q1 where xv: "x = 2 ^ (m + k) * of_nat q1" and q1v: "q1 < 2 ^ (LENGTH('a) - (m + k))" by (auto elim: is_alignedE) have l1: "2 ^ (m + k) * q1 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q1v]) (subst kv, rule order_less_imp_le [OF nlt]) have l2: "2 ^ m * q2 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q2v], rule order_less_imp_le [OF order_le_less_trans]) fact+ have "x = of_nat (2 ^ (m + k) * q1)" using xv by simp moreover have "y = of_nat (2 ^ m * q2)" using yv by simp ultimately have upls: "unat x + unat y = 2 ^ m * (2 ^ k * q1 + q2)" proof - have f1: "unat x = 2 ^ (m + k) * q1" - by (metis (no_types) \x = of_nat (2 ^ (m + k) * q1)\ l1 nat_mod_lem word_unat.inverse_norm - zero_less_numeral zero_less_power) + using q1v unat_mult_power_lem xv by blast have "unat y = 2 ^ m * q2" - by (metis (no_types) \y = of_nat (2 ^ m * q2)\ l2 nat_mod_lem word_unat.inverse_norm - zero_less_numeral zero_less_power) + using q2v unat_mult_power_lem yv by blast then show ?thesis using f1 by (simp add: power_add semiring_normalization_rules(34)) qed (* (2 ^ k * q1 + q2) *) show "\d. unat (x + y) = 2 ^ m * d" proof (cases "unat x + unat y < 2 ^ LENGTH('a)") case True have "unat (x + y) = unat x + unat y" by (subst unat_plus_if', rule if_P) fact also have "\ = 2 ^ m * (2 ^ k * q1 + q2)" by (rule upls) finally show ?thesis .. next case False then have "unat (x + y) = (unat x + unat y) mod 2 ^ LENGTH('a)" by (subst unat_word_ariths(1)) simp also have "\ = (2 ^ m * (2 ^ k * q1 + q2)) mod 2 ^ LENGTH('a)" by (subst upls, rule refl) also have "\ = 2 ^ m * ((2 ^ k * q1 + q2) mod 2 ^ (LENGTH('a) - m))" proof - have "m \ len_of (TYPE('a))" by (meson le_trans less_imp_le_nat lt nlt) then show ?thesis by (metis mult_mod_right ordered_cancel_comm_monoid_diff_class.add_diff_inverse power_add) qed finally show ?thesis .. qed qed next assume "\ n < LENGTH('a)" with assms show ?thesis by (simp add: is_aligned_mask not_less take_bit_eq_mod power_overflow word_arith_nat_defs(7) flip: take_bit_eq_mask) qed corollary aligned_sub_aligned: "\is_aligned (x::'a::len word) n; is_aligned y m; m \ n\ \ is_aligned (x - y) m" apply (simp del: add_uminus_conv_diff add:diff_conv_add_uminus) apply (erule aligned_add_aligned, simp_all) apply (erule is_alignedE) apply (rule_tac k="- of_nat q" in is_alignedI) apply simp done lemma is_aligned_shift: fixes k::"'a::len word" shows "is_aligned (push_bit m k) m" proof cases assume mv: "m < LENGTH('a)" from mv obtain q where mq: "m + q = LENGTH('a)" and "0 < q" by (auto dest: less_imp_add_positive) have "(2::nat) ^ m dvd unat (push_bit m k)" proof have kv: "(unat k div 2 ^ q) * 2 ^ q + unat k mod 2 ^ q = unat k" by (rule div_mult_mod_eq) have "unat (push_bit m k) = unat (2 ^ m * k)" - by (simp add: push_bit_eq_mult) + by (simp add: push_bit_eq_mult ac_simps) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (simp add: unat_word_ariths(2)) also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp finally show "unat (push_bit m k) = 2 ^ m * (unat k mod 2 ^ q)" . qed then show ?thesis by (unfold is_aligned_iff_dvd_nat) next assume "\ m < LENGTH('a)" then show ?thesis by (simp add: not_less power_overflow is_aligned_mask word_size) qed lemma word_mod_by_0: "k mod (0::'a::len word) = k" by (simp add: word_arith_nat_mod) lemma aligned_mod_eq_0: fixes p::"'a::len word" assumes al: "is_aligned p sz" shows "p mod 2 ^ sz = 0" proof cases assume szv: "sz < LENGTH('a)" with al show ?thesis unfolding is_aligned_iff_dvd_nat by (simp add: and_mask_dvd_nat p2_gt_0 word_mod_2p_is_mask) next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: is_aligned_mask flip: take_bit_eq_mask take_bit_eq_mod) qed lemma is_aligned_triv: "is_aligned (2 ^ n ::'a::len word) n" by (rule is_alignedI [where k = 1], simp) lemma is_aligned_mult_triv1: "is_aligned (2 ^ n * x ::'a::len word) n" by (rule is_alignedI [OF refl]) lemma is_aligned_mult_triv2: "is_aligned (x * 2 ^ n ::'a::len word) n" by (subst mult.commute, simp add: is_aligned_mult_triv1) lemma word_power_less_0_is_0: fixes x :: "'a::len word" shows "x < a ^ 0 \ x = 0" by simp lemma is_aligned_no_wrap: fixes off :: "'a::len word" fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "unat ptr + unat off < 2 ^ LENGTH('a)" proof - have szv: "sz < LENGTH('a)" using off p2_gt_0 word_neq_0_conv by fastforce from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by simp next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q ::'a::len word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply (simp_all) done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + unat off < 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_less [OF _ qv]) apply (rule order_less_le_trans [OF unat_mono [OF off] order_eq_refl]) apply simp_all done qed qed qed lemma is_aligned_no_wrap': fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "ptr \ ptr + off" by (subst no_plus_overflow_unat_size, subst word_size, rule is_aligned_no_wrap) fact+ lemma is_aligned_no_overflow': fixes p :: "'a::len word" assumes al: "is_aligned p n" shows "p \ p + (2 ^ n - 1)" proof cases assume "n n ptr \ ptr + 2^sz - 1" by (drule is_aligned_no_overflow') (simp add: field_simps) lemma replicate_not_True: "\n. xs = replicate n False \ True \ set xs" by (induct xs) auto lemma map_zip_replicate_False_xor: "n = length xs \ map (\(x, y). x = (\ y)) (zip xs (replicate n False)) = xs" by (induct xs arbitrary: n, auto) lemma drop_minus_lem: "\ n \ length xs; 0 < n; n' = length xs \ \ drop (n' - n) xs = rev xs ! (n - 1) # drop (Suc (n' - n)) xs" proof (induct xs arbitrary: n n') case Nil then show ?case by simp next case (Cons y ys) from Cons.prems show ?case apply simp apply (cases "n = Suc (length ys)") apply (simp add: nth_append) apply (simp add: Suc_diff_le Cons.hyps nth_append) apply clarsimp apply arith done qed lemma drop_minus: "\ n < length xs; n' = length xs \ \ drop (n' - Suc n) xs = rev xs ! n # drop (n' - n) xs" apply (subst drop_minus_lem) apply simp apply simp apply simp apply simp apply (cases "length xs", simp) apply (simp add: Suc_diff_le) done lemma aligned_add_xor: \(x + 2 ^ n) XOR 2 ^ n = x\ if al: \is_aligned (x::'a::len word) n'\ and le: \n < n'\ proof - have \\ bit x n\ using that by (rule is_aligned_imp_not_bit) then have \x + 2 ^ n = x OR 2 ^ n\ by (subst disjunctive_add) (auto simp add: bit_simps disjunctive_add) moreover have \(x OR 2 ^ n) XOR 2 ^ n = x\ by (rule bit_word_eqI) (auto simp add: bit_simps \\ bit x n\) ultimately show ?thesis by simp qed lemma is_aligned_add_mult_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n * z) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x*z"]) done lemma is_aligned_add_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x"]) done lemma is_aligned_no_wrap''': fixes ptr :: "'a::len word" shows"\ is_aligned ptr sz; sz < LENGTH('a); off < 2 ^ sz \ \ unat ptr + off < 2 ^ LENGTH('a)" apply (drule is_aligned_no_wrap[where off="of_nat off"]) apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod) apply (subst(asm) unat_of_nat_len) apply (erule order_less_trans) apply (erule power_strict_increasing) apply simp apply assumption done lemma is_aligned_get_word_bits: fixes p :: "'a::len word" shows "\ is_aligned p n; \ is_aligned p n; n < LENGTH('a) \ \ P; \ p = 0; n \ LENGTH('a) \ \ P \ \ P" apply (cases "n < LENGTH('a)") apply simp apply simp apply (erule meta_mp) apply (simp add: is_aligned_mask power_add power_overflow not_less flip: take_bit_eq_mask) apply (metis take_bit_length_eq take_bit_of_0 take_bit_tightened) done lemma aligned_small_is_0: "\ is_aligned x n; x < 2 ^ n \ \ x = 0" by (simp add: is_aligned_mask less_mask_eq) corollary is_aligned_less_sz: "\is_aligned a sz; a \ 0\ \ \ a < 2 ^ sz" by (rule notI, drule(1) aligned_small_is_0, erule(1) notE) lemma aligned_at_least_t2n_diff: "\is_aligned x n; is_aligned y n; x < y\ \ x \ y - 2 ^ n" apply (erule is_aligned_get_word_bits[where p=y]) apply (rule ccontr) apply (clarsimp simp: linorder_not_le) apply (subgoal_tac "y - x = 0") apply clarsimp apply (rule aligned_small_is_0) apply (erule(1) aligned_sub_aligned) apply simp apply unat_arith apply simp done lemma is_aligned_no_overflow'': "\is_aligned x n; x + 2 ^ n \ 0\ \ x \ x + 2 ^ n" apply (frule is_aligned_no_overflow') apply (erule order_trans) apply (simp add: field_simps) apply (erule word_sub_1_le) done lemma is_aligned_bitI: \is_aligned p m\ if \\n. n < m \ \ bit p n\ apply (simp add: is_aligned_mask) apply (rule bit_word_eqI) using that apply (auto simp add: bit_simps) done lemma is_aligned_nth [word_eqI_simps]: "is_aligned p m = (\n < m. \ bit p n)" apply (auto intro: is_aligned_bitI simp add: is_aligned_mask bit_eq_iff) apply (auto simp: bit_simps) using bit_imp_le_length not_less apply blast done lemma range_inter: "({a..b} \ {c..d} = {}) = (\x. \(a \ x \ x \ b \ c \ x \ x \ d))" by auto lemma aligned_inter_non_empty: "\ {p..p + (2 ^ n - 1)} \ {p..p + 2 ^ m - 1} = {}; is_aligned p n; is_aligned p m\ \ False" apply (clarsimp simp only: range_inter) apply (erule_tac x=p in allE) apply simp apply (erule impE) apply (erule is_aligned_no_overflow') apply (erule notE) apply (erule is_aligned_no_overflow) done lemma not_aligned_mod_nz: assumes al: "\ is_aligned a n" shows "a mod 2 ^ n \ 0" apply (rule ccontr) using al apply (rule notE) apply simp apply (rule is_alignedI [of _ _ \a div 2 ^ n\]) apply (metis add.right_neutral mult.commute word_mod_div_equality) done lemma nat_add_offset_le: fixes x :: nat assumes yv: "y \ 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y \ 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" by (auto simp: le_iff_add) have "x * 2 ^ n + y \ x * 2 ^ n + 2 ^ n" using yv xv by simp also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y \ 2 ^ (m + n)" . qed lemma is_aligned_no_wrap_le: fixes ptr::"'a::len word" assumes al: "is_aligned ptr sz" and szv: "sz < LENGTH('a)" and off: "off \ 2 ^ sz" shows "unat ptr + off \ 2 ^ LENGTH('a)" proof - from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by (auto simp add: le_Suc_eq Suc_le_eq) (simp add: le_less) next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q :: 'a word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply simp_all done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + off \ 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_le [OF off qv]) apply simp done qed qed qed lemma is_aligned_neg_mask: "m \ n \ is_aligned (x AND NOT (mask n)) m" by (rule is_aligned_bitI) (simp add: bit_simps) lemma unat_minus: "unat (- (x :: 'a :: len word)) = (if x = 0 then 0 else 2 ^ size x - unat x)" using unat_sub_if_size[where x="2 ^ size x" and y=x] by (simp add: unat_eq_0 word_size) lemma is_aligned_minus: \is_aligned (- p) n\ if \is_aligned p n\ for p :: \'a::len word\ using that apply (cases \n < LENGTH('a)\) apply (simp_all add: not_less is_aligned_beyond_length) apply transfer apply (simp flip: take_bit_eq_0_iff) apply (subst take_bit_minus [symmetric]) apply simp done lemma add_mask_lower_bits: "\is_aligned (x :: 'a :: len word) n; \n' \ n. n' < LENGTH('a) \ \ bit p n'\ \ x + p AND NOT (mask n) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth) apply (erule_tac x=na in allE)+ apply (simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_less word_size) apply (metis is_aligned_nth not_le) done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (push_bit m w) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (drop_bit m w) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl_self: "is_aligned (push_bit n p) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p AND NOT (mask n) = p" apply (rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) done lemma is_aligned_shiftr_shiftl: "is_aligned w n \ push_bit n (drop_bit n w) = w" apply (rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) using not_le_imp_less apply blast apply (metis add_diff_inverse_nat) done lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ push_bit n (drop_bit n x AND mask v) = x AND mask (v + n)" apply (rule word_eqI) apply (simp add: word_size bit_simps) apply (subgoal_tac "\m. bit x m \ m \ n") apply auto[1] apply (clarsimp simp: is_aligned_mask) apply (drule_tac x=m in word_eqD) apply (frule test_bit_size) apply (simp add: word_size bit_simps) done lemma mask_zero: "is_aligned x a \ x AND mask a = 0" by (metis is_aligned_mask) lemma is_aligned_neg_mask_eq_concrete: "\ is_aligned p n; msk AND NOT (mask n) = NOT (mask n) \ \ p AND msk = p" by (metis word_bw_assocs(1) word_bw_comms(1) is_aligned_neg_mask_eq) lemma is_aligned_and_not_zero: "\ is_aligned n k; n \ 0 \ \ 2 ^ k \ n" using is_aligned_less_sz leI by blast lemma is_aligned_and_2_to_k: "(n AND 2 ^ k - 1) = 0 \ is_aligned (n :: 'a :: len word) k" by (simp add: is_aligned_mask mask_eq_decr_exp) lemma is_aligned_power2: "b \ a \ is_aligned (2 ^ a) b" by (metis is_aligned_triv is_aligned_weaken) lemma aligned_sub_aligned': "\ is_aligned (a :: 'a :: len word) n; is_aligned b n; n < LENGTH('a) \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma is_aligned_neg_mask_weaken: "\ is_aligned p n; m \ n \ \ p AND NOT (mask m) = p" using is_aligned_neg_mask_eq is_aligned_weaken by blast lemma is_aligned_neg_mask2 [simp]: "is_aligned (a AND NOT (mask n)) n" by (rule is_aligned_bitI) (simp add: bit_simps) lemma is_aligned_0': "is_aligned 0 n" by (fact is_aligned_0) lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv - apply - - apply (erule is_alignedE) - apply (subst word_unat.Rep_inject [symmetric]) - apply (subst unat_mod) - apply (subst iffD1 [OF unat_add_lem], rule ux) - apply simp - apply (subst unat_mult_power_lem, assumption+) - apply (simp) - apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) - apply (erule eq_imp_le[OF unat_power_lower]) + apply (simp flip: take_bit_eq_mod) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps) + apply (metis assms(2) bit_or_iff is_aligned_mask is_aligned_nth leD less_mask_eq word_and_le1 word_bw_lcs(1) word_neq_0_conv word_plus_and_or_coroll) + apply (meson assms(2) leI less_2p_is_upper_bits_unset) + apply (metis assms(2) bit_disjunctive_add_iff bit_imp_le_length bit_push_bit_iff is_alignedE' less_2p_is_upper_bits_unset) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_eq_decr_exp word_mod_by_0) qed lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply transfer apply simp apply (auto simp add: le_less power_2_ge_iff szv) apply (metis le_less_trans mask_eq_decr_exp mask_lt_2pn order_less_imp_le szv) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_eq_decr_exp power_overflow) then show ?thesis .. qed lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d AND mask n = d) \ (p + d AND (NOT (mask n)) = p)" apply (subst (asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll) apply (simp_all flip: take_bit_eq_mask) apply (metis take_bit_eq_mask word_bw_lcs(1) word_log_esimps(1)) apply (metis add.commute add_left_imp_eq take_bit_eq_mask word_plus_and_or_coroll2) done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) AND mask n = q AND mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q AND NOT (mask n)) = (p + q) AND NOT (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p OR d" apply (subst disjunctive_add) apply (simp_all add: is_aligned_iff_take_bit_eq_0) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) subgoal for m apply (cases \m < n\) apply (auto simp add: not_less) apply (metis bit_take_bit_iff less_mask_eq take_bit_eq_mask) done done lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" by (metis NOT_mask add_diff_cancel_right' diff_0 is_aligned_neg_mask_eq not_le word_and_le1) lemma neg_mask_mono_le: "x \ y \ x AND NOT(mask n) \ y AND NOT(mask n)" for x :: "'a :: len word" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False then show "y AND NOT(mask n) < x AND NOT(mask n) \ False" by (simp add: mask_eq_decr_exp linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y AND NOT(mask n) < x AND NOT(mask n)" have word_bits: "n < LENGTH('a)" by fact have "y \ (y AND NOT(mask n)) + (y AND mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y AND NOT(mask n)) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b; simp add: is_aligned_neg_mask) done also have "\ \ x AND NOT(mask n)" using b apply (subst add.commute) apply (rule le_plus) apply (rule aligned_at_least_t2n_diff; simp add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated]; simp add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma and_neg_mask_eq_iff_not_mask_le: "w AND NOT(mask n) = NOT(mask n) \ NOT(mask n) \ w" for w :: \'a::len word\ by (metis eq_iff neg_mask_mono_le word_and_le1 word_and_le2 word_bw_same(1)) lemma neg_mask_le_high_bits [word_eqI_simps]: \NOT (mask n) \ w \ (\i \ {n ..< size w}. bit w i)\ (is \?P \ ?Q\) for w :: \'a::len word\ proof assume ?Q then have \w AND NOT (mask n) = NOT (mask n)\ by (auto simp add: bit_simps word_size intro: bit_word_eqI) then show ?P by (simp add: and_neg_mask_eq_iff_not_mask_le) next assume ?P then have *: \w AND NOT (mask n) = NOT (mask n)\ by (simp add: and_neg_mask_eq_iff_not_mask_le) show \?Q\ proof (rule ccontr) assume \\ (\i\{n.. then obtain m where m: \\ bit w m\ \n \ m\ \m < LENGTH('a)\ by (auto simp add: word_size) from * have \bit (w AND NOT (mask n)) m \ bit (NOT (mask n :: 'a word)) m\ by auto with m show False by (auto simp add: bit_simps) qed qed lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_eq_decr_exp) (erule is_aligned_no_overflow') lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: dvd_def is_aligned_iff_dvd_nat) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" by (meson is_aligned_get_word_bits) thus ?thesis by simp qed lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply (simp add: even_mask_iff) done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less not_le) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (subst (asm) unat_add_lem' [symmetric]) apply (simp add: is_aligned_mask_offset_unat) apply (metis gap_between_aligned linorder_not_less mask_eq_decr_exp unat_arith_simps(2)) done lemma aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: of_nat_diff) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" using take_bit_nat_less_self_iff [of \LENGTH('a)\ \mq - 2 ^ sq * nq\] apply (auto simp add: word_less_nat_alt not_le not_less) apply (metis take_bit_nat_eq_self_iff) done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt take_bit_eq_mod) apply (subst (asm) mod_less) apply auto apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp done ultimately show ?thesis by auto qed lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ drop_bit n (x + y) = drop_bit n y" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) apply (metis less_2p_is_upper_bits_unset not_le) apply (metis le_add1 less_2p_is_upper_bits_unset test_bit_bin) done lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ drop_bit n (y + x) = drop_bit n y" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) apply (metis less_2p_is_upper_bits_unset not_le) apply (metis bit_imp_le_length le_add1 less_2p_is_upper_bits_unset) done lemma and_neg_mask_plus_mask_mono: "(p AND NOT (mask n)) + mask n \ p" for p :: \'a::len word\ apply (rule word_le_minus_cancel[where x = "p AND NOT (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_eq_decr_exp word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr AND NOT (mask n)) + (2 ^ n - 1)" for ptr :: \'a::len word\ by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p AND mask n = d) \ (p AND (NOT (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k AND mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a AND mask n = (ptr AND mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a AND mask n) AND NOT (mask m) = (ptr + a AND NOT (mask m) ) AND mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) +lemma word_add_power_off: + fixes a :: "'a :: len word" + assumes ak: "a < k" + and kw: "k < 2 ^ (LENGTH('a) - m)" + and mw: "m < LENGTH('a)" + and off: "off < 2 ^ m" +shows "(a * 2 ^ m) + off < k * 2 ^ m" +proof (cases "m = 0") + case True + then show ?thesis using off ak by simp +next + case False + + from ak have ak1: "a + 1 \ k" by (rule inc_le) + then have "(a + 1) * 2 ^ m \ 0" + apply - + apply (rule word_power_nonzero) + apply (erule order_le_less_trans [OF _ kw]) + apply (rule mw) + apply (rule less_is_non_zero_p1 [OF ak]) + done + then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw + apply - + apply (simp add: distrib_right) + apply (rule word_plus_strict_mono_right [OF off]) + apply (rule is_aligned_no_overflow'') + apply (rule is_aligned_mult_triv2) + apply assumption + done + also have "\ \ k * 2 ^ m" using ak1 mw kw False + apply - + apply (erule word_mult_le_mono1) + apply (simp add: p2_gt_0) + apply (simp add: word_less_nat_alt) + apply (meson nat_mult_power_less_eq zero_less_numeral) + done + finally show ?thesis . +qed + +lemma offset_not_aligned: + "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ + \ is_aligned (p + of_nat i) n" + apply (erule is_aligned_add_not_aligned) + apply transfer + apply (auto simp add: is_aligned_iff_udvd) + apply (metis (no_types, lifting) le_less_trans less_not_refl2 less_or_eq_imp_le of_nat_eq_0_iff take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat unat_lt2p unat_power_lower) + done + +lemma le_or_mask: + "w \ w' \ w OR mask x \ w' OR mask x" + for w w' :: \'a::len word\ + by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) + end diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy --- a/thys/Word_Lib/Bits_Int.thy +++ b/thys/Word_Lib/Bits_Int.thy @@ -1,1454 +1,1468 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Bitwise Operations on integers\ theory Bits_Int imports "HOL-Library.Word" begin subsection \Implicit bit representation of \<^typ>\int\\ lemma bin_last_def: "(odd :: int \ bool) w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) lemma bin_last_numeral_simps [simp]: "\ odd (0 :: int)" "odd (1 :: int)" "odd (- 1 :: int)" "odd (Numeral1 :: int)" "\ odd (numeral (Num.Bit0 w) :: int)" "odd (numeral (Num.Bit1 w) :: int)" "\ odd (- numeral (Num.Bit0 w) :: int)" "odd (- numeral (Num.Bit1 w) :: int)" by simp_all lemma bin_rest_numeral_simps [simp]: "(\k::int. k div 2) 0 = 0" "(\k::int. k div 2) 1 = 0" "(\k::int. k div 2) (- 1) = - 1" "(\k::int. k div 2) Numeral1 = 0" "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd y\ \ x = y" by (auto elim: oddE) lemma [simp]: shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) x > 0 \ x > 1" by auto subsection \Bit projection\ lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" using that bin_nth_eq_iff [of x y] by (simp add: fun_eq_iff) lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" by (fact bit_eq_iff) lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) lemma bin_nth_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" by (induction n) (simp_all add: bit_Suc) lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps lemma nth_2p_bin: "(bit :: int \ nat \ bool) (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done lemma bin_nth_numeral_unfold: "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) w) = bin_sign w" by (simp add: bin_sign_def) lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) n w = w mod 2 ^ n" by (fact take_bit_eq_mod) lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: \(signed_take_bit :: nat \ int \ int) n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc) lemma sbintrunc_0_numeral [simp]: "(signed_take_bit :: nat \ int \ int) 0 1 = -1" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) n bin) = -1) = bit bin n" by (simp add: bin_sign_def) lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) w n" by (fact bit_take_bit_iff) lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by (simp add: min.absorb2) lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" by (simp add: min.absorb2) lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (n - 1) w" by (cases n) simp_all lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp moreover have \(signed_take_bit :: nat \ int \ int) n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_int_greater_eq) lemma sbintrunc_dec: \(signed_take_bit :: nat \ int \ int) n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_int_less_eq) lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) lemma bintr_lt2p: "(take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: bintrunc_mod2p) lemma bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq) lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" by (simp add: sbintrunc_mod2p) lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: sbintrunc_mod2p) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (Suc n) bin)" by (auto simp add: take_bit_Suc) lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) lemma bin_cat_eq_push_bit_add_take_bit: \concat_bit n l k = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) y n z)" by (fact concat_bit_assoc) lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: \nat \ int list \ int\ where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ lemma bin_rcat_eq_foldl: \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0 ks\ by (induction ks rule: rev_induct) (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) qed fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) k a) n b" by (metis bin_cat_assoc bin_cat_zero) lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: \drop_bit n ((\k n l. concat_bit n l k) v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: \take_bit n ((\k n l. concat_bit n l k) v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ primrec bin_sc :: "nat \ bool \ int \ int" where Z: "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" | Suc: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" lemma bin_nth_sc [bit_simps]: "bit (bin_sc n b w) n \ b" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) done lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ by (simp_all add: fun_eq_iff bit_eq_iff) (simp_all add: bin_nth_sc_gen bit_set_bit_iff bit_unset_bit_iff) lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" apply (cases x) apply (simp_all add: bin_sc_eq bit_eq_iff) apply (auto simp add: bit_take_bit_iff bit_set_bit_iff bit_unset_bit_iff) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: bin_sc_eq unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: bin_sc_eq set_bit_greater_eq) lemma bintr_bin_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" by (simp add: bin_sc_eq take_bit_unset_bit_eq unset_bit_less_eq) lemma bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" by (simp add: bin_sc_eq take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc.Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma shiftl_int_def: "push_bit n x = x * 2 ^ n" for x :: int by (fact push_bit_eq_mult) lemma shiftr_int_def: "drop_bit n x = x div 2 ^ n" for x :: int by (fact drop_bit_eq_div) subsubsection \Basic simplification rules\ lemmas int_not_def = not_int_def lemma int_not_simps [simp]: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (fact bit.conj_one_left) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (fact bit.disj_zero_left) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (fact bit.xor_zero_left) subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) y" by (subst xor_int_rec) auto lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) y)" by (subst xor_int_rec) auto lemma bin_nth_ops: "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp subsubsection \Interactions with arithmetic\ lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x OR y)" by simp_all lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) n (NOT x)" by (fact take_bit_not_take_bit) text \Want theorems of the form of \bin_trunc_xor\.\ lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ lemma not_int_cmp_0 [simp]: fixes i :: int shows "0 < NOT i \ i < -1" "0 \ NOT i \ i < 0" "NOT i < 0 \ i \ 0" "NOT i \ 0 \ i \ -1" by(simp_all add: int_not_def) arith+ lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" by (fact bit.conj_disj_distrib) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: "(odd :: int \ bool) i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0: "push_bit 0 x = x" and int_shiftl_Suc: "push_bit (Suc n) x = 2 * push_bit n x" by (auto simp add: shiftl_int_def) lemma int_0_shiftl: "push_bit n 0 = (0 :: int)" by (fact push_bit_of_0) lemma bin_last_shiftl: "odd (push_bit n x) \ n = 0 \ (odd :: int \ bool) x" by simp lemma bin_rest_shiftl: "(\k::int. k div 2) (push_bit n x) = (if n > 0 then push_bit (n - 1) x else (\k::int. k div 2) x)" by (cases n) (simp_all add: push_bit_eq_mult) lemma bin_nth_shiftl: "(bit :: int \ nat \ bool) (push_bit n x) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" by (fact bit_push_bit_iff_int) lemma bin_last_shiftr: "odd (drop_bit n x) \ bit x n" for x :: int by (simp add: bit_iff_odd_drop_bit) lemma bin_rest_shiftr: "(\k::int. k div 2) (drop_bit n x) = drop_bit (Suc n) x" by (simp add: drop_bit_Suc drop_bit_half) lemma bin_nth_shiftr: "(bit :: int \ nat \ bool) (drop_bit n x) m = (bit :: int \ nat \ bool) x (n + m)" by (simp add: bit_simps) lemma bin_nth_conv_AND: fixes x :: int shows "(bit :: int \ nat \ bool) x n \ x AND (push_bit n 1) \ 0" by (fact bit_iff_and_push_bit_not_eq_0) lemma int_shiftl_numeral [simp]: "push_bit (numeral w') (numeral w :: int) = push_bit (pred_numeral w') (numeral (num.Bit0 w))" "push_bit (numeral w') (- numeral w :: int) = push_bit (pred_numeral w') (- numeral (num.Bit0 w))" by(simp_all add: numeral_eq_Suc shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "push_bit (numeral w) (1::int) = push_bit (pred_numeral w) 2" using int_shiftl_numeral [of Num.One w] by (simp add: numeral_eq_Suc) lemma shiftl_ge_0: fixes i :: int shows "push_bit n i \ 0 \ i \ 0" by (fact push_bit_nonnegative_int_iff) lemma shiftl_lt_0: fixes i :: int shows "push_bit n i < 0 \ i < 0" by (fact push_bit_negative_int_iff) lemma int_shiftl_test_bit: "bit (push_bit i n :: int) m \ m \ i \ bit n (m - i)" by (fact bit_push_bit_iff_int) lemma int_0shiftr: "drop_bit x (0 :: int) = 0" by (fact drop_bit_of_0) lemma int_minus1_shiftr: "drop_bit x (-1 :: int) = -1" by (fact drop_bit_minus_one) lemma int_shiftr_ge_0: fixes i :: int shows "drop_bit n i \ 0 \ i \ 0" by (fact drop_bit_nonnegative_int_iff) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "drop_bit n i < 0 \ i < 0" by (fact drop_bit_negative_int_iff) lemma int_shiftr_numeral [simp]: "drop_bit (numeral w') (1 :: int) = 0" "drop_bit (numeral w') (numeral num.One :: int) = 0" "drop_bit (numeral w') (numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (- numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (- numeral w)" "drop_bit (numeral w') (- numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (- numeral (Num.inc w))" by (simp_all add: numeral_eq_Suc add_One drop_bit_Suc) lemma int_shiftr_numeral_Suc0 [simp]: "drop_bit (Suc 0) (1 :: int) = 0" "drop_bit (Suc 0) (numeral num.One :: int) = 0" "drop_bit (Suc 0) (numeral (num.Bit0 w) :: int) = numeral w" "drop_bit (Suc 0) (numeral (num.Bit1 w) :: int) = numeral w" "drop_bit (Suc 0) (- numeral (num.Bit0 w) :: int) = - numeral w" "drop_bit (Suc 0) (- numeral (num.Bit1 w) :: int) = - numeral (Num.inc w)" by (simp_all add: drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = push_bit n 1" and m: "m < n" and x: "x < y" shows "bit (x - y) m = bit x m" proof - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ by (simp_all add: bin_sign_def push_bit_eq_mult split: if_splits) from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ proof (induction m arbitrary: x n) case 0 then show ?case by simp next case (Suc m) moreover define q where \q = n - 1\ ultimately have n: \n = Suc q\ by simp have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ by simp moreover from Suc.IH [of \x div 2\ q] Suc.prems have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ by (simp add: n) ultimately show ?case by (simp add: bit_Suc n) qed with \y = 2 ^ n\ show ?thesis by simp qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) lemma bin_set_conv_OR: "bin_sc n True i = i OR (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) subsection \More lemmas on words\ lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], folded uint_word_of_int_eq, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma setBit_no: "set_bit n (numeral bin) = word_of_int (bin_sc n True (numeral bin))" by transfer (simp add: bin_sc_eq) lemma clearBit_no: "unset_bit n (numeral bin) = word_of_int (bin_sc n False (numeral bin))" by transfer (simp add: bin_sc_eq) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" by transfer (simp add: drop_bit_take_bit ac_simps) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int ((\k n l. concat_bit n l k) w (size b) (uint b))" by transfer (simp add: take_bit_concat_bit_eq) lemma bintrunc_shiftl: "take_bit n (push_bit i m) = push_bit i (take_bit (n - i) m)" for m :: int by (fact take_bit_push_bit) lemma uint_shiftl: "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))" by (simp add: unsigned_push_bit_eq word_size) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = push_bit n (1 :: int)" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma sint_range': \- (2 ^ (LENGTH('a) - Suc 0)) \ sint x \ sint x < 2 ^ (LENGTH('a) - Suc 0)\ for x :: \'a::len word\ apply transfer using sbintr_ge sbintr_lt apply auto done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self min.absorb2 simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed lemma bintrunc_id: "\m \ int n; 0 < m\ \ take_bit n m = m" by (simp add: take_bit_int_eq_self_iff le_less_trans less_exp) +lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" + if "n = m" "a = c" "take_bit m b = take_bit m d" + using that(3) unfolding that(1,2) + by (simp add: bin_cat_eq_push_bit_add_take_bit) + +lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" + by (metis drop_bit_bin_cat_eq) + +lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" + by (metis take_bit_bin_cat_eq) + +lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" + by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) + code_identifier code_module Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Word_Lib/Enumeration_Word.thy b/thys/Word_Lib/Enumeration_Word.thy --- a/thys/Word_Lib/Enumeration_Word.thy +++ b/thys/Word_Lib/Enumeration_Word.thy @@ -1,298 +1,329 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Thomas Sewell *) section "Enumeration Instances for Words" theory Enumeration_Word imports "HOL-Library.Word" More_Word Enumeration Even_More_List begin lemma length_word_enum: "length (enum :: 'a :: len word list) = 2 ^ LENGTH('a)" by (simp add: enum_word_def) lemma fromEnum_unat[simp]: "fromEnum (x :: 'a::len word) = unat x" proof - have "enum ! the_index enum x = x" by (auto intro: nth_the_index) moreover have "the_index enum x < length (enum::'a::len word list)" by (auto intro: the_index_bounded) moreover { fix y assume "of_nat y = x" moreover assume "y < 2 ^ LENGTH('a)" ultimately have "y = unat x" using of_nat_inverse by fastforce } ultimately show ?thesis by (simp add: fromEnum_def enum_word_def) qed lemma toEnum_of_nat[simp]: "n < 2 ^ LENGTH('a) \ (toEnum n :: 'a :: len word) = of_nat n" by (simp add: toEnum_def length_word_enum enum_word_def) instantiation word :: (len) enumeration_both begin definition enum_alt_word_def: "enum_alt \ alt_from_ord (enum :: ('a :: len) word list)" instance by (intro_classes, simp add: enum_alt_word_def) end definition upto_enum_step :: "('a :: len) word \ 'a word \ 'a word \ 'a word list" ("[_ , _ .e. _]") where "upto_enum_step a b c \ if c < a then [] else map (\x. a + x * (b - a)) [0 .e. (c - a) div (b - a)]" (* in the wraparound case, bad things happen. *) lemma maxBound_word: "(maxBound::'a::len word) = -1" by (simp add: maxBound_def enum_word_def last_map of_nat_diff) lemma minBound_word: "(minBound::'a::len word) = 0" by (simp add: minBound_def enum_word_def upt_conv_Cons) lemma maxBound_max_word: "(maxBound::'a::len word) = - 1" by (fact maxBound_word) lemma leq_maxBound [simp]: "(x::'a::len word) \ maxBound" by (simp add: maxBound_max_word) lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < LENGTH('a)" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "x < y \ [x::'a::len word .e. y] = x # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons) apply simp_all apply unat_arith apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply simp_all apply unat_arith done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply simp_all apply (metis le_unat_uoi nat_less_le toEnum_of_nat unsigned_less word_le_nat_alt) apply (metis le_unat_uoi less_or_eq_imp_le toEnum_of_nat unsigned_less word_le_nat_alt) apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply transfer apply auto done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv le_m1_iff_lt p2_gt_0 by auto qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_cases: fixes a :: "'a::len word" shows "length [a .e. b] = (if a \ b then Suc (unat b) - unat a else 0)" apply (case_tac "a \ b") apply (clarsimp) apply (clarsimp simp: upto_enum_def) apply unat_arith done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done +lemma remdups_enum_upto: + fixes s::"'a::len word" + shows "remdups [s .e. e] = [s .e. e]" + by simp + +lemma card_enum_upto: + fixes s::"'a::len word" + shows "card (set [s .e. e]) = Suc (unat e) - unat s" + by (subst List.card_set) (simp add: remdups_enum_upto) + +lemma length_upto_enum_one: + fixes x :: "'a :: len word" + assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" + shows "[x , y .e. z] = [x]" + unfolding upto_enum_step_def +proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) + show "unat ((z - x) div (y - x)) = 0" + proof (subst unat_div, rule div_less) + have syx: "unat (y - x) = unat y - unat x" + by (rule unat_sub [OF order_less_imp_le]) fact + moreover have "unat (z - x) = unat z - unat x" + by (rule unat_sub) fact + + ultimately show "unat (z - x) < unat (y - x)" + using lt2 lt3 unat_mono word_less_minus_mono_left by blast + qed + + then show "(z - x) div (y - x) * (y - x) = 0" + by (simp add: unat_div) (simp add: word_arith_nat_defs(6)) +qed + end diff --git a/thys/Word_Lib/More_Word.thy b/thys/Word_Lib/More_Word.thy --- a/thys/Word_Lib/More_Word.thy +++ b/thys/Word_Lib/More_Word.thy @@ -1,2168 +1,2544 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Lemmas on words\ theory More_Word imports "HOL-Library.Word" More_Arithmetic More_Divides begin +\ \problem posed by TPHOLs referee: + criterion for overflow of addition of signed integers\ + +lemma sofl_test: + \sint x + sint y = sint (x + y) \ + drop_bit (size x - 1) ((x + y XOR x) AND (x + y XOR y)) = 0\ + for x y :: \'a::len word\ +proof - + obtain n where n: \LENGTH('a) = Suc n\ + by (cases \LENGTH('a)\) simp_all + have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ + \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ + using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] + by (auto intro: ccontr) + have \sint x + sint y = sint (x + y) \ + (sint (x + y) < 0 \ sint x < 0) \ + (sint (x + y) < 0 \ sint y < 0)\ + using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] + signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] + apply (auto simp add: not_less) + apply (unfold sint_word_ariths) + apply (subst signed_take_bit_int_eq_self) + prefer 4 + apply (subst signed_take_bit_int_eq_self) + prefer 7 + apply (subst signed_take_bit_int_eq_self) + prefer 10 + apply (subst signed_take_bit_int_eq_self) + apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) + apply (smt (z3) take_bit_nonnegative) + apply (smt (z3) take_bit_int_less_exp) + apply (smt (z3) take_bit_nonnegative) + apply (smt (z3) take_bit_int_less_exp) + done + then show ?thesis + apply (simp only: One_nat_def word_size drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) + apply (simp add: bit_last_iff) + done +qed + lemma unat_power_lower [simp]: "unat ((2::'a::len word) ^ n) = 2 ^ n" if "n < LENGTH('a::len)" using that by transfer simp lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" by (fact unat_power_lower) lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by transfer simp lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply (simp only: word_arith_nat_defs word_le_nat_alt word_of_nat_eq_iff flip: nat_div_eq_Suc_0_iff) apply (simp flip: unat_div unsigned_take_bit_eq) done lemma AND_twice [simp]: "(w AND m) AND m = w AND m" by (fact and.right_idem) lemma word_combine_masks: "w AND m = z \ w AND m' = z' \ w AND (m OR m') = (z OR z')" for w m m' z z' :: \'a::len word\ by (simp add: bit.conj_disj_distrib) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" by (simp add : word_gt_0 not_le) lemma uint_2p_alt: \n < LENGTH('a::len) \ uint ((2::'a::len word) ^ n) = 2 ^ n\ using p2_gt_0 [of n, where ?'a = 'a] by (simp add: uint_2p) lemma p2_eq_0: \(2::'a::len word) ^ n = 0 \ LENGTH('a::len) \ n\ by (fact exp_eq_zero_iff) lemma p2len: \(2 :: 'a word) ^ LENGTH('a::len) = 0\ by simp lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps simp flip: push_bit_eq_mult drop_bit_eq_div) lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" for w :: \'a::len word\ by (rule neg_mask_is_div) lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size simp flip: push_bit_eq_mult drop_bit_eq_div) lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule and_mask_arith) lemma mask_2pm1: "mask n = 2 ^ n - (1 :: 'a::len word)" by (fact mask_eq_decr_exp) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" for x :: \'a::len word\ by (simp add: mask_eq_decr_exp) lemma word_and_mask_le_2pm1: "w AND mask n \ 2 ^ n - 1" for w :: \'a::len word\ by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u AND mask n = 0 \ v < 2^n \ u AND v = 0" for u v :: \'a::len word\ apply (drule less_mask_eq) apply (simp flip: take_bit_eq_mask) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) done lemma and_mask_eq_iff_le_mask: \w AND mask n = w \ w \ mask n\ for w :: \'a::len word\ apply (simp flip: take_bit_eq_mask) apply (cases \n \ LENGTH('a)\; transfer) apply (simp_all add: not_le min_def) apply (simp_all add: mask_eq_exp_minus_1) apply auto apply (metis take_bit_int_less_exp) apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) done lemma less_eq_mask_iff_take_bit_eq_self: \w \ mask n \ take_bit n w = w\ for w :: \'a::len word\ by (simp add: and_mask_eq_iff_le_mask take_bit_eq_mask) lemma NOT_eq: "NOT (x :: 'a :: len word) = - x - 1" apply (cut_tac x = "x" in word_add_not) apply (drule add.commute [THEN trans]) apply (drule eq_diff_eq [THEN iffD2]) by simp lemma NOT_mask: "NOT (mask n :: 'a::len word) = - (2 ^ n)" by (simp add : NOT_eq mask_2pm1) lemma le_m1_iff_lt: "(x > (0 :: 'a :: len word)) = ((y \ x - 1) = (y < x))" by uint_arith lemma gt0_iff_gem1: \0 < x \ x - 1 < x\ for x :: \'a::len word\ by (metis add.right_neutral diff_add_cancel less_irrefl measure_unat unat_arith_simps(2) word_neq_0_conv word_sub_less_iff) lemma power_2_ge_iff: \2 ^ n - (1 :: 'a::len word) < 2 ^ n \ n < LENGTH('a)\ using gt0_iff_gem1 p2_gt_0 by blast lemma le_mask_iff_lt_2n: "n < len_of TYPE ('a) = (((w :: 'a :: len word) \ mask n) = (w < 2 ^ n))" unfolding mask_2pm1 by (rule trans [OF p2_gt_0 [THEN sym] le_m1_iff_lt]) lemma mask_lt_2pn: \n < LENGTH('a) \ mask n < (2 :: 'a::len word) ^ n\ by (simp add: mask_eq_exp_minus_1 power_2_ge_iff) lemma word_unat_power: "(2 :: 'a :: len word) ^ n = of_nat (2 ^ n)" by simp lemma of_nat_mono_maybe: assumes xlt: "x < 2 ^ len_of TYPE ('a)" shows "y < x \ of_nat y < (of_nat x :: 'a :: len word)" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (erule order_less_trans [OF _ xlt]) apply (subst mod_less [OF xlt]) apply assumption done lemma word_and_max_word: fixes a::"'a::len word" shows "x = - 1 \ a AND x = a" by simp lemma word_and_full_mask_simp: \x AND mask LENGTH('a) = x\ for x :: \'a::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ (0 :: 'a word)\ then have \n < LENGTH('a)\ by simp then show \bit (x AND Bit_Operations.mask LENGTH('a)) n \ bit x n\ by (simp add: bit_and_iff bit_mask_iff) qed lemma of_int_uint: "of_int (uint x) = x" by (fact word_of_int_uint) corollary word_plus_and_or_coroll: "x AND y = 0 \ x + y = x OR y" for x y :: \'a::len word\ using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x AND w) + (x AND NOT w) = x" for x w :: \'a::len word\ apply (subst disjunctive_add) apply (simp add: bit_simps) apply (simp flip: bit.conj_disj_distrib) done lemma nat_mask_eq: \nat (mask n) = mask n\ by (simp add: nat_eq_iff of_nat_mask_eq) lemma unat_mask_eq: \unat (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer (simp add: nat_mask_eq) lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by transfer (simp add: min_def split: if_splits) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" apply transfer apply (cases \LENGTH('c) = LENGTH('a)\) apply (auto simp add: min_def) done lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_eq_unatI: \v = w\ if \unat v = unat w\ using that by transfer (simp add: nat_eq_iff) lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_eq_unatI) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma enum_word_nth_eq: \(Enum.enum :: 'a::len word list) ! n = word_of_nat n\ if \n < 2 ^ LENGTH('a)\ for n using that by (simp add: enum_word_def) lemma length_enum_word_eq: \length (Enum.enum :: 'a::len word list) = 2 ^ LENGTH('a)\ by (simp add: enum_word_def) lemma unat_lt2p [iff]: \unat x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply simp_all done lemma word_sub_1_le: "x \ 0 \ x - 1 \ (x :: ('a :: len) word)" apply (subst no_ulen_sub) apply simp apply (cases "uint x = 0") apply (simp add: uint_0_iff) apply (insert uint_ge_0[where x=x]) apply arith done lemma push_bit_word_eq_nonzero: \push_bit n w \ 0\ if \w < 2 ^ m\ \m + n < LENGTH('a)\ \w \ 0\ for w :: \'a::len word\ using that apply (simp only: word_neq_0_conv word_less_nat_alt mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (metis add_diff_cancel_right' gr0I gr_implies_not0 less_or_eq_imp_le min_def push_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_push_bit take_bit_take_bit unsigned_push_bit_eq) done lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof (cases \sz < LENGTH('a)\) case True with assms show ?thesis by (simp add: unat_word_ariths take_bit_eq_mod mod_simps) (simp add: take_bit_nat_eq_self_iff nat_less_power_trans flip: take_bit_eq_mod) next case False with assms show ?thesis by simp qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma word_le_minus_one_leq: "x < y \ x \ y - 1" for x :: "'a :: len word" by transfer (metis le_less_trans less_irrefl take_bit_decr_eq take_bit_nonnegative zle_diff1_eq) lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" using le_less_trans word_le_minus_one_leq nv power_2_ge_iff by blast lemma unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" by (simp add: take_bit_nat_eq_self_iff) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by transfer (auto simp add: take_bit_of_nat nat_eq_iff take_bit_nat_eq_self_iff intro: sym) lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma mask_out_sub_mask: "(x AND NOT (mask n)) = x - (x AND (mask n))" for x :: \'a::len word\ by (simp add: field_simps word_plus_and_or_coroll2) lemma subtract_mask: "p - (p AND mask n) = (p AND NOT (mask n))" "p - (p AND NOT (mask n)) = (p AND mask n)" for p :: \'a::len word\ by (simp add: field_simps word_plus_and_or_coroll2)+ lemma take_bit_word_eq_self_iff: \take_bit n w = w \ n \ LENGTH('a) \ w < 2 ^ n\ for w :: \'a::len word\ using take_bit_int_eq_self_iff [of n \take_bit LENGTH('a) (uint w)\] by (transfer fixing: n) auto lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a::len)" "y < LENGTH('a::len)" shows "x < y" using x using assms by transfer simp lemma mask_twice: "(x AND mask n) AND mask m = x AND mask (min m n)" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask) lemma plus_one_helper[elim!]: "x < n + (1 :: 'a :: len word) \ x \ n" apply (simp add: word_less_nat_alt word_le_nat_alt field_simps) apply (case_tac "1 + n = 0") apply simp_all apply (subst(asm) unatSuc, assumption) apply arith done lemma plus_one_helper2: "\ x \ n; n + 1 \ 0 \ \ x < n + (1 :: 'a :: len word)" by (simp add: word_less_nat_alt word_le_nat_alt field_simps unatSuc) lemma less_x_plus_1: fixes x :: "'a :: len word" shows "x \ - 1 \ (y < (x + 1)) = (y < x \ y = x)" apply (rule iffI) apply (rule disjCI) apply (drule plus_one_helper) apply simp apply (subgoal_tac "x < x + 1") apply (erule disjE, simp_all) apply (rule plus_one_helper2 [OF order_refl]) apply (rule notI, drule max_word_wrap) apply simp done lemma word_Suc_leq: fixes k::"'a::len word" shows "k \ - 1 \ x < k + 1 \ x \ k" using less_x_plus_1 word_le_less_eq by auto lemma word_Suc_le: fixes k::"'a::len word" shows "x \ - 1 \ x + 1 \ k \ x < k" by (meson not_less word_Suc_leq) lemma word_lessThan_Suc_atMost: \{..< k + 1} = {..k}\ if \k \ - 1\ for k :: \'a::len word\ using that by (simp add: lessThan_def atMost_def word_Suc_leq) lemma word_atLeastLessThan_Suc_atLeastAtMost: \{l ..< u + 1} = {l..u}\ if \u \ - 1\ for l :: \'a::len word\ using that by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost) lemma word_atLeastAtMost_Suc_greaterThanAtMost: \{m<..u} = {m + 1..u}\ if \m \ - 1\ for m :: \'a::len word\ using that by (simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le) lemma word_atLeastLessThan_Suc_atLeastAtMost_union: fixes l::"'a::len word" assumes "m \ - 1" and "l \ m" and "m \ u" shows "{l..m} \ {m+1..u} = {l..u}" proof - from ivl_disj_un_two(8)[OF assms(2) assms(3)] have "{l..u} = {l..m} \ {m<..u}" by blast with assms show ?thesis by(simp add: word_atLeastAtMost_Suc_greaterThanAtMost) qed lemma max_word_less_eq_iff [simp]: \- 1 \ w \ w = - 1\ for w :: \'a::len word\ by (fact word_order.extremum_unique) lemma word_or_zero: "(a OR b = 0) = (a = 0 \ b = 0)" for a b :: \'a::len word\ by (fact or_eq_0_iff) lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" shows "2^n < (2::'a::len word)^m" by (smt suc_n le_less_trans lessI nat_less_le nat_mult_less_cancel_disj p2_gt_0 power_Suc power_Suc unat_power_lower word_less_nat_alt x) lemma power_overflow: "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" by simp lemmas extra_sle_sless_unfolds [simp] = word_sle_eq[where a=0 and b=1] word_sle_eq[where a=0 and b="numeral n"] word_sle_eq[where a=1 and b=0] word_sle_eq[where a=1 and b="numeral n"] word_sle_eq[where a="numeral n" and b=0] word_sle_eq[where a="numeral n" and b=1] word_sless_alt[where a=0 and b=1] word_sless_alt[where a=0 and b="numeral n"] word_sless_alt[where a=1 and b=0] word_sless_alt[where a=1 and b="numeral n"] word_sless_alt[where a="numeral n" and b=0] word_sless_alt[where a="numeral n" and b=1] for n lemma word_sint_1: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" by transfer simp lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (signed_take_bit (LENGTH('a::len) - Suc 0) (1::int)))" by transfer simp lemma scast_1: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" apply (simp only: flip: mask_eq_exp_minus_1) apply transfer apply (simp add: take_bit_minus_one_eq_mask nat_mask_eq) done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma neg_mask_combine: "NOT(mask a) AND NOT(mask b) = NOT(mask (max a b) :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma neg_mask_twice: "x AND NOT(mask n) AND NOT(mask m) = x AND NOT(mask (max n m))" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma multiple_mask_trivia: "n \ m \ (x AND NOT(mask n)) + (x AND mask n AND NOT(mask m)) = x AND NOT(mask m)" for x :: \'a::len word\ apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_eq_decr_exp not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: bit_iff_odd) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" apply transfer apply (rule signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle numeral_power_eq_of_nat_cancel_iff) done lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le order_less_le_trans take_bit_nat_eq_self_iff) apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le take_bit_nat_eq_self_iff) apply (simp flip: of_nat_take_bit add: take_bit_nat_eq_self) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by transfer simp lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) apply (metis bit_take_bit_iff less_mask_eq not_less take_bit_eq_mask) done lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" by (fact word_le_minus_one_leq) lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: take_bit_eq_mod) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_eq_nat_uint [symmetric]) apply (subst unat_eq_nat_uint [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by transfer (simp add: take_bit_eq_mod) lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (metis unat_of_nat_len) lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x AND mask (len_of TYPE ('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" proof - have \2 ^ LENGTH('a) = nat (2 ^ LENGTH('a))\ by simp moreover have \unat (ucast x :: 'a word) = unat x mod nat (2 ^ LENGTH('a))\ by transfer (simp flip: nat_mod_distrib take_bit_eq_mod) ultimately show ?thesis by (simp only:) qed lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (fact bits_div_by_1) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" by (fact word_order.extremum_unique) lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis order_refl take_bit_signed_take_bit take_bit_tightened) done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma unat_eq_1: \unat x = Suc 0 \ x = 1\ by (auto intro!: unsigned_word_eqI [where ?'a = nat]) lemma word_unat_Rep_inject1: \unat x = unat 1 \ x = 1\ by (simp add: unat_eq_1) lemma and_not_mask_twice: "(w AND NOT (mask n)) AND NOT (mask m) = w AND NOT (mask (max m n))" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_less_cases: "x < y \ x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma mask_and_mask: "mask a AND mask b = (mask (min a b) :: 'a::len word)" by (simp flip: take_bit_eq_mask ac_simps) lemma mask_eq_0_eq_x: "(x AND w = 0) = (x AND NOT w = x)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x AND w = x) = (x AND NOT w = 0)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma compl_of_1: "NOT 1 = (-2 :: 'a :: len word)" by (fact not_one) lemma split_word_eq_on_mask: "(x = y) = (x AND m = y AND m \ x AND NOT m = y AND NOT m)" for x y m :: \'a::len word\ apply transfer apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps ac_simps) done lemma word_FF_is_mask: "0xFF = (mask 8 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma word_1FF_is_mask: "0x1FF = (mask 9 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply transfer apply (auto simp add: take_bit_of_nat min_def not_le) apply (metis linorder_not_less min_def take_bit_nat_eq_self take_bit_take_bit) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma word_div_less: "m < n \ m div n = 0" for m :: "'a :: len word" by (simp add: unat_mono word_arith_nat_defs(6)) lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by transfer simp lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp_all add: word_less_nat_alt less_le_trans take_bit_eq_mod) done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (cases \2 \ LENGTH('a)\) apply simp_all apply transfer apply auto apply (metis add.commute add.right_neutral even_two_times_div_two mod_div_trivial mod_pos_pos_trivial mult.commute mult_zero_left not_less not_take_bit_negative odd_two_times_div_two_succ) done lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power del: of_nat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w AND mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply auto using of_nat_inj apply blast done lemma mask_AND_NOT_mask: "(w AND NOT (mask n)) AND mask n = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) lemma AND_NOT_mask_plus_AND_mask_eq: "(w AND NOT (mask n)) + (w AND mask n) = w" for w :: \'a::len word\ apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x AND mask n = y AND mask n" and m2: "x AND NOT (mask n) = y AND NOT (mask n)" shows "x = y" proof - have *: \x = x AND mask n OR x AND NOT (mask n)\ for x :: \'a word\ by (rule bit_word_eqI) (auto simp add: bit_simps) from assms * [of x] * [of y] show ?thesis by simp qed lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x AND 1) = (x AND 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply transfer apply simp apply (subst (2) take_bit_add [symmetric]) apply (subst take_bit_add [symmetric]) apply simp done lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma word_0_sle_from_less: \0 \s x\ if \x < 2 ^ (LENGTH('a) - 1)\ for x :: \'a::len word\ using that apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis bit_take_bit_iff min_def nat_less_le not_less_eq take_bit_int_eq_self_iff take_bit_take_bit) done lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" apply transfer apply (subst (asm) take_bit_incr_eq) apply (auto simp add: diff_less_eq) using take_bit_int_less_exp le_less_trans by blast lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp flip: signed_take_bit_eq_iff_take_bit_eq) done lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (fact signed_eq_0_iff) (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: P if \\ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P\ \\ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P\ \\ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P\ proof (cases \LENGTH('a) = 1\) case True then have \a = 0 \ a = 1\ by transfer auto with True that show ?thesis by auto next case False with that show ?thesis by (simp add: less_le Suc_le_eq) qed lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp add: signed_take_bit_int_eq_self) done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply simp done lemma uint_range': \0 \ uint x \ uint x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" by (simp add: signed_take_bit_int_eq_self) lemma of_int_sint: "of_int (sint a) = a" by simp lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply transfer apply (simp add: signed_take_bit_take_bit) done lemma word_less_nowrapI': "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = (2 ^ n :: 'a::len word)" by (clarsimp simp: mask_eq_decr_exp) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_eq unat_eq_nat_uint apply transfer apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp only: flip: ucast_nat_def) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply simp apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp only: flip: ucast_nat_def) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply simp apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp_all add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ ucast x \ (ucast y :: 'b word)" for x y :: \'a::len word\ by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma neg_mask_add_mask: "((x:: 'a :: len word) AND NOT (mask n)) + (2 ^ n - 1) = x OR mask n" unfolding mask_2pm1 [symmetric] apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w AND mask n) AND NOT (mask n) = 0" by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_mask_iff) lemma and_and_not[simp]:"(a AND b) AND NOT b = 0" for a b :: \'a::len word\ apply (subst word_bw_assocs(1)) apply clarsimp done lemma ex_mask_1[simp]: "(\x. mask x = (1 :: 'a::len word))" apply (rule_tac x=1 in exI) apply (simp add:mask_eq_decr_exp) done lemma not_switch:"NOT a = x \ a = NOT x" by auto lemma test_bit_eq_iff: "bit u = bit v \ u = v" for u v :: "'a::len word" by (simp add: bit_eq_iff fun_eq_iff) lemma test_bit_size: "bit w n \ n < size w" for w :: "'a::len word" by transfer simp lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) for x y :: "'a::len word" by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) lemma word_eqI: "(\n. n < size u \ bit u n = bit v n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ bit u x = bit v x" for u v :: "'a::len word" by simp lemma test_bit_bin': "bit w n \ n < size w \ bit (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma word_test_bit_def: \bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bit x n) (bit :: 'a::len word \ _)" by transfer_prover lemma test_bit_wi: "bit (word_of_int x :: 'a::len word) n \ n < LENGTH('a) \ bit x n" by transfer simp lemma word_ops_nth_size: "n < size x \ bit (x OR y) n = (bit x n | bit y n) \ bit (x AND y) n = (bit x n \ bit y n) \ bit (x XOR y) n = (bit x n \ bit y n) \ bit (NOT x) n = (\ bit x n)" for x :: "'a::len word" by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) lemma word_ao_nth: "bit (x OR y) n = (bit x n | bit y n) \ bit (x AND y) n = (bit x n \ bit y n)" for x :: "'a::len word" by transfer (auto simp add: bit_or_iff bit_and_iff) lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bit (sint w) n = (if n < l - 1 then bit w n else bit w (l - 1))" unfolding sint_uint l_def by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) lemma test_bit_2p: "bit (word_of_int (2 ^ n)::'a::len word) m \ m = n \ m < LENGTH('a)" by transfer (auto simp add: bit_exp_iff) lemma nth_w2p: "bit ((2::'a::len word) ^ n) m \ m = n \ m < LENGTH('a::len)" by transfer (auto simp add: bit_exp_iff) lemma bang_is_le: "bit x m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtrans(3)) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemma test_bit_1 [iff]: "bit (1 :: 'a::len word) n \ n = 0" by transfer (auto simp add: bit_1_iff) lemma nth_0: "\ bit (0 :: 'a::len word) n" by transfer simp lemma nth_minus1: "bit (-1 :: 'a::len word) n \ n < LENGTH('a)" by transfer simp lemma nth_ucast: "bit (ucast w::'a::len word) n = (bit w n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma drop_bit_numeral_bit0_1 [simp]: \drop_bit (Suc 0) (numeral k) = (word_of_int (drop_bit (Suc 0) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by (metis Word_eq_word_of_int drop_bit_word.abs_eq of_int_numeral) lemma nth_mask: \bit (mask n :: 'a::len word) i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: word_size Word.bit_mask_iff) lemma nth_slice: "bit (slice n w :: 'a::len word) m = (bit w (m + n) \ m < LENGTH('a))" apply (auto simp add: bit_simps less_diff_conv dest: bit_imp_le_length) using bit_imp_le_length apply fastforce done lemma test_bit_cat [OF refl]: "wc = word_cat a b \ bit wc n = (n < size wc \ (if n < size b then bit b n else bit a (n - size b)))" apply (simp add: word_size not_less; transfer) apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. bit b n = (n < size b \ bit c n) \ bit a m = (m < size a \ bit c (m + size b)))" by (auto simp add: word_split_bin' bit_unsigned_iff word_size bit_drop_bit_eq ac_simps dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. bit b n \ n < size b \ bit c n) \ (\m::nat. bit a m \ m < size a \ bit c (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. bit b n = (n < size b \ bit c n)) \ (\m::nat. bit a m = (m < size a \ bit c (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ bit rc n = (n < size rc \ n div sw < size wl \ bit ((rev wl) ! (n div sw)) (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def rev_map bit_horner_sum_uint_exp_iff bit_simps not_le) lemmas test_bit_cong = arg_cong [where f = "bit", THEN fun_cong] lemma max_test_bit: "bit (- 1::'a::len word) n \ n < LENGTH('a)" by (fact nth_minus1) lemma map_nth_0 [simp]: "map (bit (0::'a::len word)) xs = replicate (length xs) False" by (simp flip: map_replicate_const) lemma word_and_1: "n AND 1 = (if bit n 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff bit_1_iff intro: gr0I) lemma test_bit_1': "bit (1 :: 'a :: len word) n \ 0 < LENGTH('a) \ n = 0" by simp lemma nth_w2p_same: "bit (2^n :: 'a :: len word) n = (n < LENGTH('a))" by (simp add: nth_w2p) lemma word_leI: "(\n. \n < size (u::'a::len word); bit u n \ \ bit (v::'a::len word) n) \ u <= v" apply (rule order_trans [of u \u AND v\ v]) apply (rule eq_refl) apply (rule bit_word_eqI) apply (auto simp add: bit_simps word_and_le1 word_size) done lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. bit x n = bit y n)" by (auto simp add: bit_eq_iff) lemma neg_mask_test_bit: "bit (NOT(mask n) :: 'a :: len word) m = (n \ m \ m < LENGTH('a))" by (auto simp add: bit_simps) lemma upper_bits_unset_is_l2p: \(\n' \ n. n' < LENGTH('a) \ \ bit p n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) if \n < LENGTH('a)\ for p :: "'a :: len word" proof assume ?Q then show ?P by (meson bang_is_le le_less_trans not_le word_power_increasing) next assume ?P have \take_bit n p = p\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ show \bit (take_bit n p) q \ bit p q\ proof (cases \q < n\) case True then show ?thesis by (auto simp add: bit_simps) next case False then have \n \ q\ by simp with \?P\ \q < LENGTH('a)\ have \\ bit p q\ by simp then show ?thesis by (simp add: bit_simps) qed qed with that show ?Q using take_bit_word_eq_self_iff [of n p] by auto qed lemma less_2p_is_upper_bits_unset: "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ bit p n')" for p :: "'a :: len word" by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) lemma test_bit_over: "n \ size (x::'a::len word) \ (bit x n) = False" by transfer auto lemma le_mask_high_bits: "w \ mask n \ (\i \ {n ..< size w}. \ bit w i)" for w :: \'a::len word\ apply (auto simp add: bit_simps word_size less_eq_mask_iff_take_bit_eq_self) apply (metis bit_take_bit_iff leD) apply (metis atLeastLessThan_iff leI take_bit_word_eq_self_iff upper_bits_unset_is_l2p) done lemma test_bit_conj_lt: "(bit x m \ m < LENGTH('a)) = bit x m" for x :: "'a :: len word" using test_bit_bin by blast lemma neg_test_bit: "bit (NOT x) n = (\ bit x n \ n < LENGTH('a))" for x :: "'a::len word" by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) lemma nth_bounded: "\bit (x :: 'a :: len word) n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (rule ccontr) apply (auto simp add: not_less) apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) done lemma and_neq_0_is_nth: \x AND y \ 0 \ bit x n\ if \y = 2 ^ n\ for x y :: \'a::len word\ apply (simp add: bit_eq_iff bit_simps) using that apply (simp add: bit_simps not_le) apply transfer apply auto done lemma nth_is_and_neq_0: "bit (x::'a::len word) n = (x AND 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma max_word_not_less [simp]: "\ - 1 < x" for x :: \'a::len word\ by (fact word_order.extremum_strict) lemma bit_twiddle_min: "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" by (rule bit_eqI) (auto simp add: bit_simps) lemma bit_twiddle_max: "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" by (rule bit_eqI) (auto simp add: bit_simps max_def) lemma swap_with_xor: "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x AND mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w AND NOT(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ bit x 0" using even_plus_one_iff [of x] by simp lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = bit x 0" by transfer (simp add: even_nat_iff) lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply auto apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply auto apply (metis unat_of_nat) done lemma lsb_this_or_next: "\ (bit ((x::('a::len) word) + 1) 0) \ bit x 0" by simp lemma mask_or_not_mask: "x AND mask n OR x AND NOT (mask n) = x" for x :: \'a::len word\ apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) +lemma revcast_down_us [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (signed_drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma revcast_down_ss [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (signed_drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma revcast_down_uu [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma revcast_down_su [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma cast_down_rev [OF refl]: + "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (push_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + done + +lemma revcast_up [OF refl]: + "rc = revcast \ source_size rc + n = target_size rc \ + rc w = push_bit n (ucast w :: 'a::len word)" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + apply auto + apply (metis add.commute add_diff_cancel_right) + apply (metis diff_add_inverse2 diff_diff_add) + done + +lemmas rc1 = revcast_up [THEN + revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] +lemmas rc2 = revcast_down_uu [THEN + revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] + +lemma word_ops_nth: + fixes x y :: \'a::len word\ + shows + word_or_nth: "bit (x OR y) n = (bit x n \ bit y n)" and + word_and_nth: "bit (x AND y) n = (bit x n \ bit y n)" and + word_xor_nth: "bit (x XOR y) n = (bit x n \ bit y n)" + by (simp_all add: bit_simps) + +lemma word_power_nonzero: + "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ + \ x * 2 ^ n \ 0" + by (metis gr_implies_not0 mult_eq_0_iff nat_mult_power_less_eq numeral_2_eq_2 + p2_gt_0 unat_eq_zero unat_less_power unat_mult_lem unat_power_lower word_gt_a_gt_0 zero_less_Suc) + +lemma less_1_helper: + "n \ m \ (n - 1 :: int) < m" + by arith + +lemma div_power_helper: + "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" + apply (simp flip: mask_eq_exp_minus_1 drop_bit_eq_div) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps not_le) + done + +lemma max_word_mask: + "(- 1 :: 'a::len word) = mask LENGTH('a)" + by (fact minus_1_eq_mask) + +lemmas mask_len_max = max_word_mask[symmetric] + +lemma mask_out_first_mask_some: + "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" + for x y :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) + +lemma mask_lower_twice: + "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" + for x :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) + +lemma mask_lower_twice2: + "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" + for a :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma ucast_and_neg_mask: + "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma ucast_and_mask: + "ucast (x AND mask n) = ucast x AND mask n" + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma ucast_mask_drop: + "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma mask_exceed: + "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" + by (rule bit_word_eqI) (simp add: bit_simps) + +lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" + using less_x_plus_1 order_less_le by blast + +lemma lt_plus_1_le_word: + fixes x :: "'a::len word" + assumes bound:"n < unat (maxBound::'a word)" + shows "x < 1 + of_nat n = (x \ of_nat n)" + by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) + +lemma unat_ucast_up_simp: + fixes x :: "'a::len word" + assumes "LENGTH('a) \ LENGTH('b)" + shows "unat (ucast x :: 'b::len word) = unat x" + apply (rule bit_eqI) + using assms apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma unat_ucast_less_no_overflow: + "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" + by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp + +lemma unat_ucast_less_no_overflow_simp: + "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" + using unat_less_helper unat_ucast_less_no_overflow by blast + +lemma unat_ucast_no_overflow_le: + assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" + and upward_cast: "LENGTH('a) < LENGTH('b)" + shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" +proof - + have LR: "ucast f < b \ unat f < unat b" + apply (rule unat_less_helper) + apply (simp add:ucast_nat_def) + apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) + apply (rule upward_cast) + apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt + unat_power_lower[OF upward_cast] no_overflow) + done + have RL: "unat f < unat b \ ucast f < b" + proof- + assume ineq: "unat f < unat b" + have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" + apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) + apply (simp only: flip: ucast_nat_def) + apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) + done + then show ?thesis + apply (rule order_less_le_trans) + apply (simp add:ucast_ucast_mask word_and_le2) + done + qed + then show ?thesis by (simp add:RL LR iffI) +qed + +lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] + +lemma minus_one_word: + "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" + by simp + +lemma le_2p_upper_bits: + "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ + \n'\n. n' < LENGTH('a) \ \ bit p n'" + by (subst upper_bits_unset_is_l2p; simp) + +lemma le2p_bits_unset: + "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ bit (p::'a::len word) n'" + using upper_bits_unset_is_l2p [where p=p] + by (cases "n < LENGTH('a)") auto + +lemma complement_nth_w2p: + shows "n' < LENGTH('a) \ bit (NOT (2 ^ n :: 'a::len word)) n' = (n' \ n)" + by (fastforce simp: word_ops_nth_size word_size nth_w2p) + +lemma word_unat_and_lt: + "unat x < n \ unat y < n \ unat (x AND y) < n" + by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) + +lemma word_unat_mask_lt: + "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" + by (rule word_unat_and_lt) (simp add: unat_mask word_size) + +lemma word_sless_sint_le:"x sint x \ sint y - 1" + by (metis word_sless_alt zle_diff1_eq) + +lemma upper_trivial: + fixes x :: "'a::len word" + shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" + by (simp add: less_le) + +lemma constraint_expand: + fixes x :: "'a::len word" + shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" + by (rule mem_Collect_eq) + +lemma card_map_elide: + "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" +proof - + let ?of_nat = "of_nat :: nat \ 'a word" + have "inj_on ?of_nat {i. i < CARD('a word)}" + by (rule inj_onI) (simp add: card_word of_nat_inj) + moreover have "{0.. {i. i < CARD('a word)}" + using that by auto + ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. LENGTH('a) \ x = ucast y \ ucast x = y" + for x :: "'a::len word" and y :: "'b::len word" + by transfer simp + +lemma le_ucast_ucast_le: + "x \ ucast y \ ucast x \ y" + for x :: "'a::len word" and y :: "'b::len word" + by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) + +lemma less_ucast_ucast_less: + "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" + for x :: "'a::len word" and y :: "'b::len word" + by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) + +lemma ucast_le_ucast: + "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" + for x :: "'a::len word" + by (simp add: unat_arith_simps(1) unat_ucast_up_simp) + +lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] + +lemma ucast_or_distrib: + fixes x :: "'a::len word" + fixes y :: "'a::len word" + shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" + by (fact unsigned_or_eq) + +lemma word_exists_nth: + "(w::'a::len word) \ 0 \ \i. bit w i" + by (simp add: bit_eq_iff) + +lemma max_word_not_0 [simp]: + "- 1 \ (0 :: 'a::len word)" + by simp + +lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" + using unat_gt_0 [of \- 1 :: 'a::len word\] by simp + + +(* Miscellaneous conditional injectivity rules. *) + +lemma mult_pow2_inj: + assumes ws: "m + n \ LENGTH('a)" + assumes le: "x \ mask m" "y \ mask m" + assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" + shows "x = y" +proof (rule bit_word_eqI) + fix q + assume \q < LENGTH('a)\ + from eq have \push_bit n x = push_bit n y\ + by (simp add: push_bit_eq_mult) + moreover from le have \take_bit m x = x\ \take_bit m y = y\ + by (simp_all add: less_eq_mask_iff_take_bit_eq_self) + ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ + by simp_all + with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ + apply (simp add: push_bit_take_bit) + unfolding bit_eq_iff + apply (simp add: bit_simps not_le) + apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) + done +qed + +lemma word_of_nat_inj: + assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" + assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" + shows "x = y" + by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") + (auto dest: bounded[THEN of_nat_mono_maybe]) + +lemma word_of_int_bin_cat_eq_iff: + "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) = + word_of_int (concat_bit LENGTH('b) (uint d) (uint c)) \ b = d \ a = c" + if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" + for a::"'a::len word" and b::"'b::len word" +proof - + from that show ?thesis + using that concat_bit_eq_iff [of \LENGTH('b)\ \uint b\ \uint a\ \uint d\ \uint c\] + apply (simp add: word_of_int_eq_iff take_bit_int_eq_self flip: word_eq_iff_unsigned) + apply (simp add: concat_bit_def take_bit_int_eq_self bintr_uint take_bit_push_bit) + done +qed + +lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" + if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" + for a::"'a::len word" and b::"'b::len word" + using word_of_int_bin_cat_eq_iff [OF that, of b a d c] + by (simp add: word_cat_eq' ac_simps) + +lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" +proof - + have "2 ^ n = (1::'a word) \ n = 0" + by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 + power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) + then show ?thesis by auto +qed + end diff --git a/thys/Word_Lib/Typedef_Morphisms.thy b/thys/Word_Lib/Typedef_Morphisms.thy --- a/thys/Word_Lib/Typedef_Morphisms.thy +++ b/thys/Word_Lib/Typedef_Morphisms.thy @@ -1,368 +1,400 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson and Gerwin Klein, NICTA Consequences of type definition theorems, and of extended type definition. *) section \Type Definition Theorems\ theory Typedef_Morphisms imports Main "HOL-Library.Word" Bit_Comprehension Bits_Int begin subsection "More lemmas about normal type definitions" lemma tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" and tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" and tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" by (auto simp: type_definition_def) lemma td_nat_int: "type_definition int nat (Collect ((\) 0))" unfolding type_definition_def by auto context type_definition begin declare Rep [iff] Rep_inverse [simp] Rep_inject [simp] lemma Abs_eqD: "Abs x = Abs y \ x \ A \ y \ A \ x = y" by (simp add: Abs_inject) lemma Abs_inverse': "r \ A \ Abs r = a \ Rep a = r" by (safe elim!: Abs_inverse) lemma Rep_comp_inverse: "Rep \ f = g \ Abs \ g = f" using Rep_inverse by auto lemma Rep_eqD [elim!]: "Rep x = Rep y \ x = y" by simp lemma Rep_inverse': "Rep a = r \ Abs r = a" by (safe intro!: Rep_inverse) lemma comp_Abs_inverse: "f \ Abs = g \ g \ Rep = f" using Rep_inverse by auto lemma set_Rep: "A = range Rep" proof (rule set_eqI) show "x \ A \ x \ range Rep" for x by (auto dest: Abs_inverse [of x, symmetric]) qed lemma set_Rep_Abs: "A = range (Rep \ Abs)" proof (rule set_eqI) show "x \ A \ x \ range (Rep \ Abs)" for x by (auto dest: Abs_inverse [of x, symmetric]) qed lemma Abs_inj_on: "inj_on Abs A" unfolding inj_on_def by (auto dest: Abs_inject [THEN iffD1]) lemma image: "Abs ` A = UNIV" by (fact Abs_image) lemmas td_thm = type_definition_axioms lemma fns1: "Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr \ Abs \ fr \ Rep = fa" by (auto dest: Rep_comp_inverse elim: comp_Abs_inverse simp: o_assoc) lemmas fns1a = disjI1 [THEN fns1] lemmas fns1b = disjI2 [THEN fns1] lemma fns4: "Rep \ fa \ Abs = fr \ Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr" by auto end interpretation nat_int: type_definition int nat "Collect ((\) 0)" by (rule td_nat_int) declare nat_int.Rep_cases [cases del] nat_int.Abs_cases [cases del] nat_int.Rep_induct [induct del] nat_int.Abs_induct [induct del] subsection "Extended form of type definition predicate" lemma td_conds: "norm \ norm = norm \ fr \ norm = norm \ fr \ norm \ fr \ norm = fr \ norm \ norm \ fr \ norm = norm \ fr" apply safe apply (simp_all add: comp_assoc) apply (simp_all add: o_assoc) done lemma fn_comm_power: "fa \ tr = tr \ fr \ fa ^^ n \ tr = tr \ fr ^^ n" apply (rule ext) apply (induct n) apply (auto dest: fun_cong) done lemmas fn_comm_power' = ext [THEN fn_comm_power, THEN fun_cong, unfolded o_def] locale td_ext = type_definition + fixes norm assumes eq_norm: "\x. Rep (Abs x) = norm x" begin lemma Abs_norm [simp]: "Abs (norm x) = Abs x" using eq_norm [of x] by (auto elim: Rep_inverse') lemma td_th: "g \ Abs = f \ f (Rep x) = g x" by (drule comp_Abs_inverse [symmetric]) simp lemma eq_norm': "Rep \ Abs = norm" by (auto simp: eq_norm) lemma norm_Rep [simp]: "norm (Rep x) = Rep x" by (auto simp: eq_norm' intro: td_th) lemmas td = td_thm lemma set_iff_norm: "w \ A \ w = norm w" by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) lemma inverse_norm: "Abs n = w \ Rep w = norm n" apply (rule iffI) apply (clarsimp simp add: eq_norm) apply (simp add: eq_norm' [symmetric]) done lemma norm_eq_iff: "norm x = norm y \ Abs x = Abs y" by (simp add: eq_norm' [symmetric]) lemma norm_comps: "Abs \ norm = Abs" "norm \ Rep = Rep" "norm \ norm = norm" by (auto simp: eq_norm' [symmetric] o_def) lemmas norm_norm [simp] = norm_comps lemma fns5: "Rep \ fa \ Abs = fr \ fr \ norm = fr \ norm \ fr = fr" by (fold eq_norm') auto text \ following give conditions for converses to \td_fns1\ \<^item> the condition \norm \ fr \ norm = fr \ norm\ says that \fr\ takes normalised arguments to normalised results \<^item> \norm \ fr \ norm = norm \ fr\ says that \fr\ takes norm-equivalent arguments to norm-equivalent results \<^item> \fr \ norm = fr\ says that \fr\ takes norm-equivalent arguments to the same result \<^item> \norm \ fr = fr\ says that \fr\ takes any argument to a normalised result \ lemma fns2: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = fr \ norm \ Rep \ fa = fr \ Rep" apply (fold eq_norm') apply safe prefer 2 apply (simp add: o_assoc) apply (rule ext) apply (drule_tac x="Rep x" in fun_cong) apply auto done lemma fns3: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr" apply (fold eq_norm') apply safe prefer 2 apply (simp add: comp_assoc) apply (rule ext) apply (drule_tac f="a \ b" for a b in fun_cong) apply simp done lemma fns: "fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr \ Rep \ fa = fr \ Rep" apply safe apply (frule fns1b) prefer 2 apply (frule fns1a) apply (rule fns3 [THEN iffD1]) prefer 3 apply (rule fns2 [THEN iffD1]) apply (simp_all add: comp_assoc) apply (simp_all add: o_assoc) done lemma range_norm: "range (Rep \ Abs) = A" by (simp add: set_Rep_Abs) end lemmas td_ext_def' = td_ext_def [unfolded type_definition_def td_ext_axioms_def] subsection \Type-definition locale instantiations\ definition uints :: "nat \ int set" \ \the sets of integers representing the words\ where "uints n = range (take_bit n)" definition sints :: "nat \ int set" where "sints n = range (signed_take_bit (n - 1))" lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" by (simp add: uints_def range_bintrunc) lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" by (simp add: sints_def range_sbintrunc) definition unats :: "nat \ nat set" where "unats n = {i. i < 2 ^ n}" \ \naturals\ lemma uints_unats: "uints n = int ` unats n" apply (unfold unats_def uints_num) apply safe apply (rule_tac image_eqI) apply (erule_tac nat_0_le [symmetric]) by auto lemma unats_uints: "unats n = nat ` uints n" by (auto simp: uints_unats image_iff) lemma td_ext_uint: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) (\w::int. w mod 2 ^ LENGTH('a))" apply (unfold td_ext_def') apply transfer apply (simp add: uints_num take_bit_eq_mod) done interpretation word_uint: td_ext "uint::'a::len word \ int" word_of_int "uints (LENGTH('a::len))" "\w. w mod 2 ^ LENGTH('a::len)" by (fact td_ext_uint) lemmas td_uint = word_uint.td_thm lemmas int_word_uint = word_uint.eq_norm lemma td_ext_ubin: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) (take_bit (LENGTH('a)))" apply standard apply transfer apply simp done interpretation word_ubin: td_ext "uint::'a::len word \ int" word_of_int "uints (LENGTH('a::len))" "take_bit (LENGTH('a::len))" by (fact td_ext_ubin) lemma td_ext_unat [OF refl]: "n = LENGTH('a::len) \ td_ext (unat :: 'a word \ nat) of_nat (unats n) (\i. i mod 2 ^ n)" apply (standard; transfer) apply (simp_all add: unats_def take_bit_of_nat take_bit_nat_eq_self_iff flip: take_bit_eq_mod) done lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] interpretation word_unat: td_ext "unat::'a::len word \ nat" of_nat "unats (LENGTH('a::len))" "\i. i mod 2 ^ LENGTH('a::len)" by (rule td_ext_unat) lemmas td_unat = word_unat.td_thm lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" for z :: "'a::len word" apply (unfold unats_def) apply clarsimp apply (metis le_unat_uoi unsigned_less) done lemma td_ext_sbin: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (signed_take_bit (LENGTH('a) - 1))" by (standard; transfer) (auto simp add: sints_def) lemma td_ext_sint: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (\w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1))" using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) text \ We do \sint\ before \sbin\, before \sint\ is the user version and interpretations do not produce thm duplicates. I.e. we get the name \word_sint.Rep_eqD\, but not \word_sbin.Req_eqD\, because the latter is the same thm as the former. \ interpretation word_sint: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "\w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - 2 ^ (LENGTH('a::len) - 1)" by (rule td_ext_sint) interpretation word_sbin: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "signed_take_bit (LENGTH('a::len) - 1)" by (rule td_ext_sbin) lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] lemmas td_sint = word_sint.td lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" by (fact uints_def [unfolded no_bintr_alt1]) lemmas bintr_num = word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemmas sbintr_num = word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] interpretation test_bit: td_ext "bit :: 'a::len word \ nat \ bool" set_bits "{f. \i. f i \ i < LENGTH('a::len)}" "(\h i. h i \ i < LENGTH('a::len))" by standard (auto simp add: bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) lemmas td_nth = test_bit.td_thm lemma sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done +lemma uints_mono_iff: "uints l \ uints m \ l \ m" + using power_increasing_iff[of "2::int" l m] + apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) + apply (meson less_irrefl not_le zero_le_numeral zero_le_power) + done + +lemmas uints_monoI = uints_mono_iff[THEN iffD2] + +lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" + using that + by (auto simp: uints_num) + +lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" + using Bit_in_uints_Suc[OF that(1)] that(2) + by auto + +lemma bin_cat_in_uintsI: + \concat_bit n b a \ uints m\ if \a \ uints l\ \m \ l + n\ +proof - + from \m \ l + n\ obtain q where \m = l + n + q\ + using le_Suc_ex by blast + then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ + by (simp add: ac_simps power_add) + moreover have \a mod 2 ^ (l + q) = a\ + using \a \ uints l\ + by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) + ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ + by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) + then show ?thesis + by (simp add: uints_def) +qed + end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,2459 +1,1942 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Signed_Division_Word Signed_Words More_Word Most_significant_bit Enumeration_Word Aligned Bit_Shifts_Infix_Syntax begin +lemma ucast_le_ucast_eq: + fixes x y :: "'a::len word" + assumes x: "x < 2 ^ n" + assumes y: "y < 2 ^ n" + assumes n: "n = LENGTH('b::len)" + shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" + apply (rule iffI) + apply (cases "LENGTH('b) < LENGTH('a)") + apply (subst less_mask_eq[OF x, symmetric]) + apply (subst less_mask_eq[OF y, symmetric]) + apply (unfold n) + apply (subst ucast_ucast_mask[symmetric])+ + apply (simp add: ucast_le_ucast)+ + apply (erule ucast_mono_le[OF _ y[unfolded n]]) + done + +lemma ucast_zero_is_aligned: + \is_aligned w n\ if \UCAST('a::len \ 'b::len) w = 0\ \n \ LENGTH('b)\ +proof (rule is_aligned_bitI) + fix q + assume \q < n\ + moreover have \bit (UCAST('a::len \ 'b::len) w) q = bit 0 q\ + using that by simp + with \q < n\ \n \ LENGTH('b)\ show \\ bit w q\ + by (simp add: bit_simps) +qed + +lemma unat_ucast_eq_unat_and_mask: + "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" + apply (simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: ac_simps) + done + +lemma le_max_word_ucast_id: + \UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ + if \x \ UCAST('b::len \ 'a) (- 1)\ + for x :: \'a::len word\ +proof - + from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ + by simp + have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ + \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ + (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = + - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ + \ (1::int) \ 2 ^ LENGTH('b) \ + 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" + by force + have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" + using mod_pos_pos_trivial by force + have "(1::int) \ 2 ^ LENGTH('b)" + by simp + then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" + using f3 f2 by blast + then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" + by linarith + have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" + using a1 by force + have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" + by force + have f7: "- (1::int) * 1 = - 1" + by auto + have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" + by force + then have "x \ 2 ^ LENGTH('b) - 1" + using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) + then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ + by (simp add: word_le_def) + then have \uint x \ 2 ^ LENGTH('b) - 1\ + by (simp add: uint_word_ariths) + (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) + then show ?thesis + apply (simp add: unsigned_ucast_eq take_bit_word_eq_self_iff) + apply (meson \x \ 2 ^ LENGTH('b) - 1\ not_le word_less_sub_le) + done +qed + lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) lemma bit_shiftl_word_iff [bit_simps]: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: bit_push_bit_iff not_le) lemma shiftl_def: \w << n = ((*) 2 ^^ n) w\ for w :: \'a::len word\ proof - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) then show ?thesis by transfer simp qed lemma shiftr_def: \w >> n = ((\w. w div 2) ^^ n) w\ for w :: \'a::len word\ proof - have \(\w. w div 2) ^^ n = (drop_bit n :: 'a word \ 'a word)\ by (induction n) (simp_all add: drop_bit_half drop_bit_Suc) then show ?thesis by simp qed lemma bit_shiftr_word_iff: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: bit_simps) lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (signed_drop_bit (Suc 0) ^^ n) w\ apply (rule sym) apply (induction n) apply simp_all done lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div) lemma sshiftr_0: "0 >>> n = 0" by (fact signed_drop_bit_of_0) lemma sshiftr_n1: "-1 >>> n = -1" by (fact signed_drop_bit_of_minus_1) lemma bit_sshiftr_word_iff: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ by (fact bit_signed_drop_bit_iff) lemma nth_sshiftr : "bit (w >>> m) n = (n < size w \ (if n + m \ size w then bit w (size w - 1) else bit w (n + m)))" apply (auto simp add: bit_simps word_size ac_simps not_less) apply (meson bit_imp_le_length bit_shiftr_word_iff leD) done -lemma sshiftr_numeral [simp]: +lemma sshiftr_numeral: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ - apply (rule bit_word_eqI) - apply (simp add: word_size nth_sshiftr ac_simps bit_simps) - done - -lemma revcast_down_us [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_ucast_iff bit_sshiftr_word_iff ac_simps) - done - -lemma revcast_down_ss [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_scast_iff bit_sshiftr_word_iff ac_simps) - done + by (fact signed_drop_bit_word_numeral) lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" using sint_signed_drop_bit_eq [of n w] by (simp add: drop_bit_eq_div) lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ - by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) + by (simp add: mask_eq_exp_minus_1 push_bit_of_1) lemma shiftl_0: "(0::'a::len word) << n = 0" by (fact push_bit_of_0) lemma shiftr_0: "(0::'a::len word) >> n = 0" by (fact drop_bit_of_0) lemma nth_shiftl': "bit (w << m) n \ n < size w \ n >= m \ bit w (n - m)" for w :: "'a::len word" by transfer (auto simp add: bit_push_bit_iff) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr: "bit (w >> m) n = bit w (n + m)" for w :: "'a::len word" by (simp add: bit_simps ac_simps) lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" by (fact uint_shiftr_eq) lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) +lemmas ucast_up = + rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] +lemmas ucast_down = + rc2 [simplified rev_shiftr revcast_ucast [symmetric]] + lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (simp add: push_bit_eq_mult) +lemma word_shift_by_2: + "x * 4 = (x::'a::len word) << 2" + by (simp add: shiftl_t2n) + lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done -lemma revcast_down_uu [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_ucast_iff bit_shiftr_word_iff ac_simps) - done - -lemma revcast_down_su [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_scast_iff bit_shiftr_word_iff ac_simps) - done - -lemma cast_down_rev [OF refl]: - "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) - done - -lemma revcast_up [OF refl]: - "rc = revcast \ source_size rc + n = target_size rc \ - rc w = (ucast w :: 'a::len word) << n" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) - apply auto - apply (metis add.commute add_diff_cancel_right) - apply (metis diff_add_inverse2 diff_diff_add) - done - -lemmas rc1 = revcast_up [THEN - revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] -lemmas rc2 = revcast_down_uu [THEN - revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] - -lemmas ucast_up = - rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] -lemmas ucast_down = - rc2 [simplified rev_shiftr revcast_ucast [symmetric]] - -\ \problem posed by TPHOLs referee: - criterion for overflow of addition of signed integers\ - -lemma sofl_test: - \sint x + sint y = sint (x + y) \ - (x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\ - for x y :: \'a::len word\ -proof - - obtain n where n: \LENGTH('a) = Suc n\ - by (cases \LENGTH('a)\) simp_all - have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ - \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ - using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] - by (auto intro: ccontr) - have \sint x + sint y = sint (x + y) \ - (sint (x + y) < 0 \ sint x < 0) \ - (sint (x + y) < 0 \ sint y < 0)\ - using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] - signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] - apply (auto simp add: not_less) - apply (unfold sint_word_ariths) - apply (subst signed_take_bit_int_eq_self) - prefer 4 - apply (subst signed_take_bit_int_eq_self) - prefer 7 - apply (subst signed_take_bit_int_eq_self) - prefer 10 - apply (subst signed_take_bit_int_eq_self) - apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) - apply (smt (z3) take_bit_nonnegative) - apply (smt (z3) take_bit_int_less_exp) - apply (smt (z3) take_bit_nonnegative) - apply (smt (z3) take_bit_int_less_exp) - done - then show ?thesis - apply (simp only: One_nat_def word_size drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) - apply (simp add: bit_last_iff) - done -qed - lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) lemma shiftr_x_0: "x >> 0 = x" for x :: "'a::len word" by simp lemma shiftl_x_0: "x << 0 = x" for x :: "'a::len word" by simp lemma shiftl_1: "(1::'a::len word) << n = 2^n" by (fact push_bit_of_1) lemma shiftr_1: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) -lemma word_ops_nth: - fixes x y :: \'a::len word\ - shows - word_or_nth: "bit (x OR y) n = (bit x n \ bit y n)" and - word_and_nth: "bit (x AND y) n = (bit x n \ bit y n)" and - word_xor_nth: "bit (x XOR y) n = (bit x n \ bit y n)" - by (simp_all add: bit_simps) - lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add: uint_2p_alt word_size) apply (metis uint_shiftr_eq word_of_int_uint) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_eq_div zdiv_mono1) done +lemma le_shiftr': + "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" + apply (metis le_cases le_shiftr verit_la_disequality) + done + lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size bit_simps) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply transfer apply (auto simp add: fun_eq_iff bit_simps) apply (metis add_diff_inverse_nat) done lemma mask_shiftl_decompose: "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" by (fact push_bit_and) lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" by (fact drop_bit_and) lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" by (fact push_bit_or) lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" by (fact drop_bit_or) lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth bit_simps) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma mask_shift: "(x AND NOT (mask y)) >> y = x >> y" for x :: \'a::len word\ apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_eq_unatI) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by transfer (simp add: take_bit_push_bit) lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') apply transfer apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) done lemma shiftr_less_t2n': "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit ac_simps) done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (rule bit_word_eqI) (simp add: bit_simps) lemma signed_shift_guard_to_word: \unat x * 2 ^ y < 2 ^ n \ x = 0 \ x < 1 << n >> y\ if \n < LENGTH('a)\ \0 < n\ for x :: \'a::len word\ proof (cases \x = 0\) case True then show ?thesis by simp next case False then have \unat x \ 0\ by (simp add: unat_eq_0) then have \unat x \ 1\ by simp show ?thesis proof (cases \y < n\) case False then have \n \ y\ by simp then obtain q where \y = n + q\ using le_Suc_ex by blast moreover have \(2 :: nat) ^ n >> n + q \ 1\ by (simp add: drop_bit_eq_div power_add) ultimately show ?thesis using \x \ 0\ \unat x \ 1\ \n < LENGTH('a)\ by (simp add: power_add not_less word_le_nat_alt unat_drop_bit_eq push_bit_of_1) next case True with that have \y < LENGTH('a)\ by simp show ?thesis proof (cases \2 ^ n = unat x * 2 ^ y\) case True moreover have \unat x * 2 ^ y < 2 ^ LENGTH('a)\ using \n < LENGTH('a)\ by (simp flip: True) moreover have \(word_of_nat (2 ^ n) :: 'a word) = word_of_nat (unat x * 2 ^ y)\ using True by simp then have \2 ^ n = x * 2 ^ y\ by simp ultimately show ?thesis using \y < LENGTH('a)\ by (auto simp add: push_bit_of_1 drop_bit_eq_div word_less_nat_alt unat_div unat_word_ariths) next case False with \y < n\ have *: \unat x \ 2 ^ n div 2 ^ y\ by (auto simp flip: power_sub power_add) have \unat x * 2 ^ y < 2 ^ n \ unat x * 2 ^ y \ 2 ^ n\ using False by (simp add: less_le) also have \\ \ unat x \ 2 ^ n div 2 ^ y\ by (simp add: less_eq_div_iff_mult_less_eq) also have \\ \ unat x < 2 ^ n div 2 ^ y\ using * by (simp add: less_le) finally show ?thesis using that \x \ 0\ by (simp flip: push_bit_eq_mult drop_bit_eq_div add: push_bit_of_1 unat_drop_bit_eq word_less_iff_unsigned [where ?'a = nat]) qed qed qed lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" by (rule bit_word_eqI) (auto simp add: bit_simps word_size dest: bit_imp_le_length) lemma shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask add: take_bit_push_bit) lemma rshift_sub_mask_eq: "(a >> (size a - b)) AND mask b = a >> (size a - b)" for a :: \'a::len word\ using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" for a :: \'a::len word\ apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" for w :: \'a::len word\ by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma word_and_1_shiftl: "x AND (1 << n) = (if bit x n then (1 << n) else 0)" for x :: "'a :: len word" apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x AND (mask n << m) = ((x >> m) AND mask n) << m" for x :: \'a::len word\ apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma of_bool_nth: "of_bool (bit x v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit word_and_1) lemma shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit) done lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (bit x n))" by (simp add: and_exp_eq_0_iff_not_bit push_bit_of_1) lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" for w :: \'a::len word\ by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_push_bit_iff) (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" for x :: \'a::len word\ by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" for a b :: \'a::len word\ apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma shiftr1_unfold: "x div 2 = x >> 1" by (simp add: drop_bit_eq_div) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by (simp add: drop_bit_eq_div) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma shiftr1_irrelevant_lsb: "bit (x::('a::len) word) 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb': "\ (bit (x::('a::len) word) 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) OR (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done -lemma word_sless_sint_le:"x sint x \ sint y - 1" - by (metis word_sless_alt zle_diff1_eq) - -lemma upper_trivial: - fixes x :: "'a::len word" - shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" - by (simp add: less_le) - -lemma constraint_expand: - fixes x :: "'a::len word" - shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" - by (rule mem_Collect_eq) - -lemma card_map_elide: - "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" -proof - - let ?of_nat = "of_nat :: nat \ 'a word" - from word_unat.Abs_inj_on - have "inj_on ?of_nat {i. i < CARD('a word)}" - by (simp add: unats_def card_word) - moreover have "{0.. {i. i < CARD('a word)}" - using that by auto - ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0..UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ - if \x \ UCAST('b::len \ 'a) (- 1)\ - for x :: \'a::len word\ -proof - - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ - by simp - have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ - \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ - (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ - \ (1::int) \ 2 ^ LENGTH('b) \ - 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" - by force - have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" - using mod_pos_pos_trivial by force - have "(1::int) \ 2 ^ LENGTH('b)" - by simp - then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" - using f3 f2 by blast - then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" - by linarith - have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" - using a1 by force - have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" - by force - have f7: "- (1::int) * 1 = - 1" - by auto - have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" - by force - then have "x \ 2 ^ LENGTH('b) - 1" - using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) - then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ - by (simp add: word_le_def) - then have \uint x \ 2 ^ LENGTH('b) - 1\ - by (simp add: uint_word_ariths) - (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) - then show ?thesis - apply (simp add: word_ubin.eq_norm bintrunc_mod2p unsigned_ucast_eq) - apply (metis \x \ 2 ^ LENGTH('b) - 1\ and_mask_eq_iff_le_mask mask_eq_decr_exp take_bit_eq_mask) - done -qed - -lemma remdups_enum_upto: - fixes s::"'a::len word" - shows "remdups [s .e. e] = [s .e. e]" - by simp - -lemma card_enum_upto: - fixes s::"'a::len word" - shows "card (set [s .e. e]) = Suc (unat e) - unat s" - by (subst List.card_set) (simp add: remdups_enum_upto) - -lemma complement_nth_w2p: - shows "n' < LENGTH('a) \ bit (NOT (2 ^ n :: 'a::len word)) n' = (n' \ n)" - by (fastforce simp: word_ops_nth_size word_size nth_w2p) - -lemma word_unat_and_lt: - "unat x < n \ unat y < n \ unat (x AND y) < n" - by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) - -lemma word_unat_mask_lt: - "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" - by (rule word_unat_and_lt) (simp add: unat_mask word_size) - lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) -lemma le_or_mask: - "w \ w' \ w OR mask x \ w' OR mask x" - for w w' :: \'a::len word\ - by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) - -lemma le_shiftr': - "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" - apply (metis le_cases le_shiftr verit_la_disequality) - done - -lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" - using less_x_plus_1 order_less_le by blast - -lemma lt_plus_1_le_word: - fixes x :: "'a::len word" - assumes bound:"n < unat (maxBound::'a word)" - shows "x < 1 + of_nat n = (x \ of_nat n)" - by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) - -lemma unat_ucast_up_simp: - fixes x :: "'a::len word" - assumes "LENGTH('a) \ LENGTH('b)" - shows "unat (ucast x :: 'b::len word) = unat x" - unfolding ucast_eq unat_eq_nat_uint - apply (subst int_word_uint) - apply (subst mod_pos_pos_trivial; simp?) - apply (rule lt2p_lem) - apply (simp add: assms) - done - -lemma unat_ucast_less_no_overflow: - "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" - by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp - -lemma unat_ucast_less_no_overflow_simp: - "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" - using unat_less_helper unat_ucast_less_no_overflow by blast - -lemma unat_ucast_no_overflow_le: - assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" - and upward_cast: "LENGTH('a) < LENGTH('b)" - shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" -proof - - have LR: "ucast f < b \ unat f < unat b" - apply (rule unat_less_helper) - apply (simp add:ucast_nat_def) - apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) - apply (rule upward_cast) - apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt - unat_power_lower[OF upward_cast] no_overflow) - done - have RL: "unat f < unat b \ ucast f < b" - proof- - assume ineq: "unat f < unat b" - have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" - apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) - apply (simp only: flip: ucast_nat_def) - apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) - done - then show ?thesis - apply (rule order_less_le_trans) - apply (simp add:ucast_ucast_mask word_and_le2) - done - qed - then show ?thesis by (simp add:RL LR iffI) -qed - -lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] - -lemma minus_one_word: - "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" - by simp - -lemma mask_exceed: - "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" - by (simp add: and_not_mask) - -lemma word_shift_by_2: - "x * 4 = (x::'a::len word) << 2" - by (simp add: shiftl_t2n) - -lemma le_2p_upper_bits: - "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ - \n'\n. n' < LENGTH('a) \ \ bit p n'" - by (subst upper_bits_unset_is_l2p; simp) - -lemma le2p_bits_unset: - "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ bit (p::'a::len word) n'" - using upper_bits_unset_is_l2p [where p=p] - by (cases "n < LENGTH('a)") auto - lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done -lemma word_power_nonzero: - "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ - \ x * 2 ^ n \ 0" - by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) - shiftl_shiftr_id shiftl_t2n) - -lemma less_1_helper: - "n \ m \ (n - 1 :: int) < m" - by arith - -lemma div_power_helper: - "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" - apply (rule word_uint.Rep_eqD) - apply (simp only: uint_word_ariths uint_div uint_power_lower) - apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ - apply (subst mod_pos_pos_trivial) - apply (simp add: le_diff_eq uint_2p_alt) - apply (rule less_1_helper) - apply (rule power_increasing; simp) - apply (subst mod_pos_pos_trivial) - apply (simp add: uint_2p_alt) - apply (rule less_1_helper) - apply (rule power_increasing; simp) - apply (subst int_div_sub_1; simp add: uint_2p_alt) - apply (subst power_0[symmetric]) - apply (simp add: uint_2p_alt le_imp_power_dvd power_diff_power_eq) - done - -lemma word_add_power_off: - fixes a :: "'a :: len word" - assumes ak: "a < k" - and kw: "k < 2 ^ (LENGTH('a) - m)" - and mw: "m < LENGTH('a)" - and off: "off < 2 ^ m" - shows "(a * 2 ^ m) + off < k * 2 ^ m" -proof (cases "m = 0") - case True - then show ?thesis using off ak by simp -next - case False - - from ak have ak1: "a + 1 \ k" by (rule inc_le) - then have "(a + 1) * 2 ^ m \ 0" - apply - - apply (rule word_power_nonzero) - apply (erule order_le_less_trans [OF _ kw]) - apply (rule mw) - apply (rule less_is_non_zero_p1 [OF ak]) - done - then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw - apply - - apply (simp add: distrib_right) - apply (rule word_plus_strict_mono_right [OF off]) - apply (rule is_aligned_no_overflow'') - apply (rule is_aligned_mult_triv2) - apply assumption - done - also have "\ \ k * 2 ^ m" using ak1 mw kw False - apply - - apply (erule word_mult_le_mono1) - apply (simp add: p2_gt_0) - apply (simp add: word_less_nat_alt) - apply (meson nat_mult_power_less_eq zero_less_numeral) - done - finally show ?thesis . -qed - -lemma offset_not_aligned: - "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ - \ is_aligned (p + of_nat i) n" - apply (erule is_aligned_add_not_aligned) - apply transfer - apply (auto simp add: is_aligned_iff_udvd) - apply (metis bintrunc_bintrunc_ge int_ops(1) nat_int_comparison(1) nat_less_le take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat) - done - -lemma length_upto_enum_one: - fixes x :: "'a :: len word" - assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" - shows "[x , y .e. z] = [x]" - unfolding upto_enum_step_def -proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) - show "unat ((z - x) div (y - x)) = 0" - proof (subst unat_div, rule div_less) - have syx: "unat (y - x) = unat y - unat x" - by (rule unat_sub [OF order_less_imp_le]) fact - moreover have "unat (z - x) = unat z - unat x" - by (rule unat_sub) fact - - ultimately show "unat (z - x) < unat (y - x)" - using lt2 lt3 unat_mono word_less_minus_mono_left by blast - qed - - then show "(z - x) div (y - x) * (y - x) = 0" - by (metis mult_zero_left unat_0 word_unat.Rep_eqD) -qed - -lemma max_word_mask: - "(- 1 :: 'a::len word) = mask LENGTH('a)" - by (fact minus_1_eq_mask) - -lemmas mask_len_max = max_word_mask[symmetric] - -lemma mask_out_first_mask_some: - "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" - for x y :: \'a::len word\ - by (rule bit_word_eqI) (auto simp add: bit_simps word_size) - -lemma mask_lower_twice: - "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" - for x :: \'a::len word\ - by (rule bit_word_eqI) (auto simp add: bit_simps word_size) - -lemma mask_lower_twice2: - "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" - for a :: \'a::len word\ - by (rule bit_word_eqI) (auto simp add: bit_simps) - -lemma ucast_and_neg_mask: - "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps dest: bit_imp_le_length) - done - -lemma ucast_and_mask: - "ucast (x AND mask n) = ucast x AND mask n" - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps dest: bit_imp_le_length) - done - -lemma ucast_mask_drop: - "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps dest: bit_imp_le_length) - done - (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) (auto simp add: word_size bit_not_iff bit_push_bit_iff bit_mask_iff) (* Comparisons between different word sizes. *) -lemma eq_ucast_ucast_eq: - "LENGTH('b) \ LENGTH('a) \ x = ucast y \ ucast x = y" - for x :: "'a::len word" and y :: "'b::len word" - by transfer simp - -lemma le_ucast_ucast_le: - "x \ ucast y \ ucast x \ y" - for x :: "'a::len word" and y :: "'b::len word" - by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) - -lemma less_ucast_ucast_less: - "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" - for x :: "'a::len word" and y :: "'b::len word" - by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) - -lemma ucast_le_ucast: - "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" - for x :: "'a::len word" - by (simp add: unat_arith_simps(1) unat_ucast_up_simp) - -lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] - -lemma ucast_le_ucast_eq: - fixes x y :: "'a::len word" - assumes x: "x < 2 ^ n" - assumes y: "y < 2 ^ n" - assumes n: "n = LENGTH('b::len)" - shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" - apply (rule iffI) - apply (cases "LENGTH('b) < LENGTH('a)") - apply (subst less_mask_eq[OF x, symmetric]) - apply (subst less_mask_eq[OF y, symmetric]) - apply (unfold n) - apply (subst ucast_ucast_mask[symmetric])+ - apply (simp add: ucast_le_ucast)+ - apply (erule ucast_mono_le[OF _ y[unfolded n]]) - done - -lemma ucast_or_distrib: - fixes x :: "'a::len word" - fixes y :: "'a::len word" - shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" - by transfer simp - lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto -lemma word_exists_nth: - "(w::'a::len word) \ 0 \ \i. bit w i" - by (simp add: bit_eq_iff) - lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" - by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') + apply (auto simp add: take_bit_word_eq_self_iff word_less_nat_alt + simp flip: take_bit_eq_self_iff_drop_bit_eq_0) + apply (rule ccontr) + apply (simp add: not_le) + done lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) -lemma max_word_not_0 [simp]: - "- 1 \ (0 :: 'a::len word)" - by simp - -lemma ucast_zero_is_aligned: - \is_aligned w n\ if \UCAST('a::len \ 'b::len) w = 0\ \n \ LENGTH('b)\ -proof (rule is_aligned_bitI) - fix q - assume \q < n\ - moreover have \bit (UCAST('a::len \ 'b::len) w) q = bit 0 q\ - using that by simp - with \q < n\ \n \ LENGTH('b)\ show \\ bit w q\ - by (simp add: bit_simps) -qed - -lemma unat_ucast_eq_unat_and_mask: - "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" - apply (simp flip: take_bit_eq_mask) - apply transfer - apply (simp add: ac_simps) - done - -lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" - using unat_gt_0 [of \- 1 :: 'a::len word\] by simp - - -(* Miscellaneous conditional injectivity rules. *) - -lemma mult_pow2_inj: - assumes ws: "m + n \ LENGTH('a)" - assumes le: "x \ mask m" "y \ mask m" - assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" - shows "x = y" -proof (rule bit_word_eqI) - fix q - assume \q < LENGTH('a)\ - from eq have \push_bit n x = push_bit n y\ - by (simp add: push_bit_eq_mult) - moreover from le have \take_bit m x = x\ \take_bit m y = y\ - by (simp_all add: less_eq_mask_iff_take_bit_eq_self) - ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ - by simp_all - with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ - apply (simp add: push_bit_take_bit) - unfolding bit_eq_iff - apply (simp add: bit_simps not_le) - apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) - done -qed - -lemma word_of_nat_inj: - assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" - assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" - shows "x = y" - by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") - (auto dest: bounded[THEN of_nat_mono_maybe]) - -(* Uints *) - -lemma uints_mono_iff: "uints l \ uints m \ l \ m" - using power_increasing_iff[of "2::int" l m] - apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) - apply (meson less_irrefl not_le zero_le_numeral zero_le_power) - done - -lemmas uints_monoI = uints_mono_iff[THEN iffD2] - -lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" - using that - by (auto simp: uints_num) - -lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" - using Bit_in_uints_Suc[OF that(1)] that(2) - by auto - -lemma bin_cat_in_uintsI: - \concat_bit n b a \ uints m\ if \a \ uints l\ \m \ l + n\ -proof - - from \m \ l + n\ obtain q where \m = l + n + q\ - using le_Suc_ex by blast - then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ - by (simp add: ac_simps power_add) - moreover have \a mod 2 ^ (l + q) = a\ - using \a \ uints l\ - by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) - ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ - by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) - then show ?thesis - by (simp add: uints_def) -qed - -lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" - if "n = m" "a = c" "take_bit m b = take_bit m d" - using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) - -lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" - by (metis drop_bit_bin_cat_eq) - -lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" - by (metis take_bit_bin_cat_eq) - -lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" - by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) - -lemma word_of_int_bin_cat_eq_iff: - "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) = - word_of_int (concat_bit LENGTH('b) (uint d) (uint c)) \ b = d \ a = c" - if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" - for a::"'a::len word" and b::"'b::len word" - by (subst word_uint.Abs_inject) - (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) - -lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" - if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" - for a::"'a::len word" and b::"'b::len word" - using word_of_int_bin_cat_eq_iff [OF that, of b a d c] - by transfer auto - -lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" -proof - - have "2 ^ n = (1::'a word) \ n = 0" - by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 - power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) - then show ?thesis by auto -qed +(* continue sorting out from here *) (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) - lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a AND b = 0" by (metis and_zero_eq is_aligned_mask le_mask_imp_and_mask word_bw_lcs(1)) lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a OR b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a OR b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" apply transfer apply (auto simp add: min_def) apply (metis bintrunc_bintrunc_ge bintrunc_n_0 nat_less_le not_le take_bit_eq_0_iff) done lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" apply (simp add: less_eq_mask_iff_take_bit_eq_self) apply transfer apply (simp add: ac_simps) done lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: \p' = p \ off' = off\ if *: \p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n)\ and \is_aligned p n'\ \is_aligned p' n'\ \n' = n + LENGTH('a)\ \n' < LENGTH('b)\ proof - from \n' = n + LENGTH('a)\ have [simp]: \n' - n = LENGTH('a)\ \n + LENGTH('a) = n'\ by simp_all from \is_aligned p n'\ obtain q where p: \p = push_bit n' (word_of_nat q)\ \q < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') from \is_aligned p' n'\ obtain q' where p': \p' = push_bit n' (word_of_nat q')\ \q' < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') define m :: nat where \m = unat off\ then have off: \off = word_of_nat m\ by simp define m' :: nat where \m' = unat off'\ then have off': \off' = word_of_nat m'\ by simp have \push_bit n' q + take_bit n' (push_bit n m) < 2 ^ LENGTH('b)\ by (metis id_apply is_aligned_no_wrap''' of_nat_eq_id of_nat_push_bit p(1) p(2) take_bit_nat_eq_self_iff take_bit_nat_less_exp take_bit_push_bit that(2) that(5) unsigned_of_nat) moreover have \push_bit n' q' + take_bit n' (push_bit n m') < 2 ^ LENGTH('b)\ by (metis \n' - n = LENGTH('a)\ id_apply is_aligned_no_wrap''' m'_def of_nat_eq_id of_nat_push_bit off' p'(1) p'(2) take_bit_nat_eq_self_iff take_bit_push_bit that(3) that(5) unsigned_of_nat) ultimately have \push_bit n' q + take_bit n' (push_bit n m) = push_bit n' q' + take_bit n' (push_bit n m')\ using * by (simp add: p p' off off' push_bit_of_nat push_bit_take_bit word_of_nat_inj flip: of_nat_add) then have \int (push_bit n' q + take_bit n' (push_bit n m)) = int (push_bit n' q' + take_bit n' (push_bit n m'))\ by simp then have \concat_bit n' (int (push_bit n m)) (int q) = concat_bit n' (int (push_bit n m')) (int q')\ by (simp add: of_nat_push_bit of_nat_take_bit bin_cat_eq_push_bit_add_take_bit) then show ?thesis by (simp add: bin_cat_inj p p' off off' take_bit_of_nat take_bit_push_bit word_of_nat_eq_iff) (simp add: push_bit_eq_mult) qed lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ by (metis (no_types, hide_lams) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: of_nat_diff unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y AND mask n = 0 \ x + y AND NOT(mask n) = (x AND NOT(mask n)) + y" for x y :: \'a::len word\ by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: - "(x :: 'a :: len word) >> a << a >> a = x >> a" + "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x AND mask n = 0; y AND mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y AND mask n = mask n \ (x AND y) AND mask n = x AND mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y AND mask n = 0 \ (x AND y) AND mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) AND b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) AND b < c" by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by (drule sym) (simp flip: take_bit_eq_mask add: unsigned_ucast_eq) lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x AND mask n = 0; m \ n \ \ x AND mask m = 0" for x :: \'a::len word\ by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) AND mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) AND 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ bit p n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr ac_simps) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" apply (simp add: word_le_def) apply (simp only: uint_nat zle_int) apply transfer apply (simp add: take_bit_nat_eq_self) done lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x AND NOT(mask sz)) + 2 ^ sz - 1" for x :: \'a::len word\ by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) AND (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) AND mask m) = unat (x AND mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit) apply (rule order_trans) defer apply assumption apply (simp add: nat_le_iff of_nat_diff) done lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n AND msk = mask n \ \ (x mod m) AND msk = x mod m" for x :: \'a::len word\ by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x AND mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m AND NOT(mask n :: 'a::len word)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr AND NOT(mask sz')) - (ptr AND NOT(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") for ptr :: \'a::len word\ proof - assume lt: "sz' \ sz" hence "?lhs = ptr AND (mask sz AND NOT(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) AND NOT(mask sz) = 0" by (simp add: of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr AND NOT(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) AND NOT(mask sz)) + unat (ptr AND mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" - by (metis (no_types) of_nat_power diff_le_self le_less_trans shiftl_less_t2n - unat_less_power word_unat.Rep_inverse) + by (metis More_Word.of_nat_power nat_mult_power_less_eq numeral_2_eq_2 of_nat_push_bit push_bit_eq_mult unat_less_power unat_of_nat_len unsigned_less word_eq_unatI zero_less_Suc) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d AND mask n) = unat d \ unat (p + d AND NOT(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) AND NOT(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) AND mask high_bits = x >> low_bits" for x :: \'a::len word\ by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x AND mask low_bits = 0) = (x = 0)" for x :: \'a::len word\ using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) AND mask n = q AND mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) AND mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x AND NOT(mask sz)) + (x AND mask sz) \ (x AND NOT(mask sz)) + ((x AND mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "NOT(mask n :: 'a::len word) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * NOT(mask n) = - (x AND NOT(mask n))" for x :: \'a::len word\ by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) AND NOT(mask n) = p + (x AND NOT(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) AND NOT(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" by (rule word_eqI) (auto simp add: bit_simps) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p AND mask a >> b) :: 'a :: len word) = p AND mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p AND mask a >> b) :: 'a :: len word) = unat (p AND mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p AND mask a) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p AND mask a << c) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p AND NOT(mask a)) AND mask b = 0" for p :: \'a::len word\ using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p AND NOT(mask a)) + ((p AND mask a) AND NOT(mask b)) + (p AND mask b) = p" for p :: \'a::len word\ by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p AND mask a >> b << b) = (p AND mask a) AND NOT(mask b)" for p :: \'a::len word\ by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" - by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) + apply (simp add: push_bit_of_1 flip: push_bit_eq_mult) + apply (subst disjunctive_add) + apply (auto simp add: bit_simps) + apply (smt (z3) AND_NOT_mask_plus_AND_mask_eq and.comm_neutral and.right_idem and_not_mask bit.conj_disj_distrib bit.disj_cancel_right mask_out_first_mask_some word_bw_assocs(1) word_bw_comms(1) word_bw_comms(2) word_eq_unatI) + done lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by transfer simp lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x AND mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" - by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add - semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) + by (smt (z3) eq_or_less_helperD le_add2 le_eq_less_or_eq le_trans power_add unat_mult_lem unat_pow_le_intro unat_power_lower word_eq_unatI) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w AND mask (size w - n)" for w :: \'a::len word\ by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp flip: mul_not_mask_eq_neg_shiftl minus_exp_eq_not_mask add: push_bit_of_1) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" by transfer (simp flip: take_bit_eq_mask add: ac_simps) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "signed_take_bit n (uint (ucast w :: 'b word)) = signed_take_bit n (uint w)" by (rule bit_eqI) (use assms in \simp add: bit_simps\) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit n (uint w)) :: 'a word) i = (if n < i then bit w n else bit w i)" using assms by (simp add: bit_simps) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) i = (if LENGTH('b::len) \ i then bit w (LENGTH('b) - 1) else bit w i)" using len_a by (auto simp add: sbintrunc_uint_ucast bit_simps) lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. bit w i = bit w (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ NOT(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" apply (simp flip: take_bit_eq_mask) apply transfer apply simp done lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" by word_eqI lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done - lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') - apply (subst (1 2) int_word_uint) - apply (subst word_ubin.norm_eq_iff [symmetric]) - apply (subst (1 2) bintrunc_mod2p) - apply (insert is_down) - apply (unfold is_down_def) - apply (clarsimp simp: target_size source_size) - apply (clarsimp simp: mod_exp_eq min_def) - apply (rule distrib [symmetric]) + apply (metis local.distrib local.is_down take_bit_eq_mod ucast_down_wi uint_word_of_int_eq word_of_int_uint) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) - have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" + let ?range = \{- (2 ^ (size a - 1))..<2 ^ (size a - 1)} :: int set\ + + have result_range: "sint a sdiv sint b \ ?range \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] - apply (auto simp: max_def abs_if word_size sints_num) + apply (auto simp: max_def abs_if word_size) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: signed_divide_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] - apply (insert word_sint.Rep [where x="a"])[1] - apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) - apply (metis minus_minus sint_int_min word_sint.Rep_inject) + apply auto + apply (cases \size a\) + apply simp_all + apply (smt (z3) One_nat_def diff_Suc_1 signed_word_eqI sint_int_min sint_range' wsst_TYs(3)) done - have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" + have result_range_simple: "(sint a sdiv sint b \ ?range) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) - apply (clarsimp simp: word_size sints_num sint_int_min) + apply (clarsimp simp: word_size sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ end