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,349 @@ (* 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]) + 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 + 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) 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]) + 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) + 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) + 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) + 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] + 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) + 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,1023 @@ (* 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 (max_word::'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 (max_word::'a word) \ exponent a \ unat (max_word::'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) + 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 ) + 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) 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: nat_add_distrib nat_mult_distrib nat_power_eq + 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 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) + 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 + 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) + 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/IP_Addresses/IP_Address.thy b/thys/IP_Addresses/IP_Address.thy --- a/thys/IP_Addresses/IP_Address.thy +++ b/thys/IP_Addresses/IP_Address.thy @@ -1,414 +1,414 @@ (* Title: IP_Address.thy Authors: Cornelius Diekmann *) theory IP_Address imports "Word_Lib.Word_Lemmas" "Word_Lib.Word_Syntax" Hs_Compat WordInterval begin section \Modelling IP Adresses\ text\An IP address is basically an unsigned integer. We model IP addresses of arbitrary lengths. We will write @{typ "'i::len word"} for IP addresses of length @{term "LENGTH('i::len)"}. We use the convention to write @{typ 'i} whenever we mean IP addresses instead of generic words. When we will later have theorems with several polymorphic types in it (e.g. arbitrarily extensible packets), this notation makes it easier to spot that type @{typ 'i} is for IP addresses. The files @{file \IPv4.thy\} @{file \IPv6.thy\} concrete this for IPv4 and IPv6.\ text\The maximum IP address\ definition max_ip_addr :: "'i::len word" where "max_ip_addr \ of_nat ((2^(len_of(TYPE('i)))) - 1)" lemma max_ip_addr_max_word: "max_ip_addr = max_word" - by(simp add: max_ip_addr_def word_of_int_minus) + by (simp only: max_ip_addr_def of_nat_mask_eq flip: mask_eq_exp_minus_1) simp lemma max_ip_addr_max: "\a. a \ max_ip_addr" by(simp add: max_ip_addr_max_word) lemma range_0_max_UNIV: "UNIV = {0 .. max_ip_addr}" (*not in the simp set, for a reason*) by(simp add: max_ip_addr_max_word) fastforce lemma "size (x::'i::len word) = len_of(TYPE('i))" by(simp add:word_size) subsection\Sets of IP Addresses\ (*Warning, not executable!*) text\Specifying sets with network masks: 192.168.0.0 255.255.255.0\ definition ipset_from_netmask::"'i::len word \ 'i::len word \ 'i::len word set" where "ipset_from_netmask addr netmask \ let network_prefix = (addr AND netmask) in {network_prefix .. network_prefix OR (NOT netmask)}" text\Example (pseudo syntax): @{const ipset_from_netmask} \192.168.1.129 255.255.255.0\ = \{192.168.1.0 .. 192.168.1.255}\\ text\A network mask of all ones (i.e. @{term "(- 1)::'i::len word"}).\ lemma ipset_from_netmask_minusone: "ipset_from_netmask ip (- 1) = {ip}" by (simp add: ipset_from_netmask_def) lemma ipset_from_netmask_maxword: "ipset_from_netmask ip max_word = {ip}" by (simp add: ipset_from_netmask_def) lemma ipset_from_netmask_zero: "ipset_from_netmask ip 0 = UNIV" by (auto simp add: ipset_from_netmask_def) text\Specifying sets in Classless Inter-domain Routing (CIDR) notation: 192.168.0.0/24\ definition ipset_from_cidr ::"'i::len word \ nat \ 'i::len word set" where "ipset_from_cidr addr pflength \ ipset_from_netmask addr ((mask pflength) << (len_of(TYPE('i)) - pflength))" text\Example (pseudo syntax): @{const ipset_from_cidr} \192.168.1.129 24\ = \{192.168.1.0 .. 192.168.1.255}\\ (*does this simplify stuff?*) lemma "(case ipcidr of (base, len) \ ipset_from_cidr base len) = uncurry ipset_from_cidr ipcidr" by(simp add: uncurry_case_stmt) lemma ipset_from_cidr_0: "ipset_from_cidr ip 0 = UNIV" by(auto simp add: ipset_from_cidr_def ipset_from_netmask_def Let_def) text\A prefix length of word size gives back the singleton set with the IP address. Example: \192.168.1.2/32 = {192.168.1.2}\\ lemma ipset_from_cidr_wordlength: fixes ip :: "'i::len word" shows "ipset_from_cidr ip (LENGTH('i)) = {ip}" by (simp add: ipset_from_cidr_def ipset_from_netmask_def) text\Alternative definition: Considering words as bit lists:\ lemma ipset_from_cidr_bl: fixes addr :: "'i::len word" shows "ipset_from_cidr addr pflength \ ipset_from_netmask addr (of_bl ((replicate pflength True) @ (replicate ((len_of(TYPE('i))) - pflength)) False))" by(simp add: ipset_from_cidr_def mask_bl shiftl_of_bl) lemma ipset_from_cidr_alt: fixes pre :: "'i::len word" shows "ipset_from_cidr pre len = {pre AND (mask len << LENGTH('i) - len) .. pre OR mask (LENGTH('i) - len)}" apply(simp add: ipset_from_cidr_def ipset_from_netmask_def Let_def) apply(simp add: word_oa_dist) apply(simp add: NOT_mask_shifted_lenword) done lemma ipset_from_cidr_alt2: fixes base ::"'i::len word" shows "ipset_from_cidr base len = ipset_from_netmask base (NOT (mask (LENGTH('i) - len)))" apply(simp add: ipset_from_cidr_def) using NOT_mask_shifted_lenword by(metis word_not_not) text\In CIDR notation, we cannot express the empty set.\ lemma ipset_from_cidr_not_empty: "ipset_from_cidr base len \ {}" by(simp add: ipset_from_cidr_alt bitmagic_zeroLast_leq_or1Last) text\Though we can write 192.168.1.2/24, we say that 192.168.0.0/24 is well-formed.\ lemma ipset_from_cidr_base_wellforemd: fixes base:: "'i::len word" assumes "mask (LENGTH('i) - l) AND base = 0" shows "ipset_from_cidr base l = {base .. base OR mask (LENGTH('i) - l)}" proof - have maskshift_eq_not_mask_generic: "((mask l << LENGTH('i) - l) :: 'i::len word) = NOT (mask (LENGTH('i) - l))" using NOT_mask_shifted_lenword by (metis word_not_not) have *: "base AND NOT (mask (LENGTH('i) - l)) = base" unfolding mask_eq_0_eq_x[symmetric] using assms word_bw_comms(1)[of base] by simp hence **: "base AND NOT (mask (LENGTH('i) - l)) OR mask (LENGTH('i) - l) = base OR mask (LENGTH('i) - l)" by simp have "ipset_from_netmask base (NOT (mask (LENGTH('i) - l))) = {base .. base || mask (LENGTH('i) - l)}" by(simp add: ipset_from_netmask_def Let_def ** *) thus ?thesis by(simp add: ipset_from_cidr_def maskshift_eq_not_mask_generic) qed lemma ipset_from_cidr_large_pfxlen: fixes ip:: "'i::len word" assumes "n \ LENGTH('i)" shows "ipset_from_cidr ip n = {ip}" proof - have obviously: "mask (LENGTH('i) - n) = 0" by (simp add: assms) show ?thesis apply(subst ipset_from_cidr_base_wellforemd) subgoal using assms by simp by (simp add: obviously) qed lemma ipset_from_netmask_base_mask_consume: fixes base :: "'i::len word" shows "ipset_from_netmask (base AND NOT (mask (LENGTH('i) - m))) (NOT (mask (LENGTH('i) - m))) = ipset_from_netmask base (NOT (mask (LENGTH('i) - m)))" unfolding ipset_from_netmask_def by(simp) text\Another definition of CIDR notation: All IP address which are equal on the first @{term "len - n"} bits\ definition ip_cidr_set :: "'i::len word \ nat \ 'i word set" where "ip_cidr_set i r \ {j . i AND NOT (mask (LENGTH('i) - r)) = j AND NOT (mask (LENGTH('i) - r))}" text\The definitions are equal\ lemma ipset_from_cidr_eq_ip_cidr_set: fixes base::"'i::len word" shows "ipset_from_cidr base len = ip_cidr_set base len" proof - have maskshift_eq_not_mask_generic: "((mask len << LENGTH('a) - len) :: 'a::len word) = NOT (mask (LENGTH('a) - len))" using NOT_mask_shifted_lenword by (metis word_not_not) have 1: "mask (len - m) AND base AND NOT (mask (len - m)) = 0" for len m and base::"'i::len word" by(simp add: word_bw_lcs) have 2: "mask (LENGTH('i) - len) AND pfxm_p = 0 \ (a \ ipset_from_netmask pfxm_p (NOT (mask (LENGTH('i) - len)))) \ (pfxm_p = NOT (mask (LENGTH('i) - len)) AND a)" for a::"'i::len word" and pfxm_p apply(subst ipset_from_cidr_alt2[symmetric]) apply(subst zero_base_lsb_imp_set_eq_as_bit_operation) apply(simp; fail) apply(subst ipset_from_cidr_base_wellforemd) apply(simp; fail) apply(simp) done from 2[OF 1, of _ base] have "(x \ ipset_from_netmask base (~~ (mask (LENGTH('i) - len)))) \ (base && ~~ (mask (LENGTH('i) - len)) = x && ~~ (mask (LENGTH('i) - len)))" for x apply(simp add: ipset_from_netmask_base_mask_consume) unfolding word_bw_comms(1)[of _ " ~~ (mask (LENGTH('i) - len))"] by simp then show ?thesis unfolding ip_cidr_set_def ipset_from_cidr_def by(auto simp add: maskshift_eq_not_mask_generic) qed lemma ip_cidr_set_change_base: "j \ ip_cidr_set i r \ ip_cidr_set j r = ip_cidr_set i r" by (auto simp: ip_cidr_set_def) subsection\IP Addresses as WordIntervals\ text\The nice thing is: @{typ "'i wordinterval"}s are executable.\ definition iprange_single :: "'i::len word \ 'i wordinterval" where "iprange_single ip \ WordInterval ip ip" fun iprange_interval :: "('i::len word \ 'i::len word) \ 'i wordinterval" where "iprange_interval (ip_start, ip_end) = WordInterval ip_start ip_end" declare iprange_interval.simps[simp del] lemma iprange_interval_uncurry: "iprange_interval ipcidr = uncurry WordInterval ipcidr" by(cases ipcidr) (simp add: iprange_interval.simps) lemma "wordinterval_to_set (iprange_single ip) = {ip}" by(simp add: iprange_single_def) lemma "wordinterval_to_set (iprange_interval (ip1, ip2)) = {ip1 .. ip2}" by(simp add: iprange_interval.simps) text\Now we can use the set operations on @{typ "'i::len wordinterval"}s\ term wordinterval_to_set term wordinterval_element term wordinterval_union term wordinterval_empty term wordinterval_setminus term wordinterval_UNIV term wordinterval_invert term wordinterval_intersection term wordinterval_subset term wordinterval_eq subsection\IP Addresses in CIDR Notation\ text\We want to convert IP addresses in CIDR notation to intervals. We already have @{const ipset_from_cidr}, which gives back a non-executable set. We want to convert to something we can store in an @{typ "'i wordinterval"}.\ fun ipcidr_to_interval_start :: "('i::len word \ nat) \ 'i::len word" where "ipcidr_to_interval_start (pre, len) = ( let netmask = (mask len) << (LENGTH('i) - len); network_prefix = (pre AND netmask) in network_prefix)" fun ipcidr_to_interval_end :: "('i::len word \ nat) \ 'i::len word" where "ipcidr_to_interval_end (pre, len) = ( let netmask = (mask len) << (LENGTH('i) - len); network_prefix = (pre AND netmask) in network_prefix OR (NOT netmask))" definition ipcidr_to_interval :: "('i::len word \ nat) \ ('i word \ 'i word)" where "ipcidr_to_interval cidr \ (ipcidr_to_interval_start cidr, ipcidr_to_interval_end cidr)" lemma ipset_from_cidr_ipcidr_to_interval: "ipset_from_cidr base len = {ipcidr_to_interval_start (base,len) .. ipcidr_to_interval_end (base,len)}" by(simp add: Let_def ipcidr_to_interval_def ipset_from_cidr_def ipset_from_netmask_def) declare ipcidr_to_interval_start.simps[simp del] ipcidr_to_interval_end.simps[simp del] lemma ipcidr_to_interval: "ipcidr_to_interval (base, len) = (s,e) \ ipset_from_cidr base len = {s .. e}" by (simp add: ipcidr_to_interval_def ipset_from_cidr_ipcidr_to_interval) definition ipcidr_tuple_to_wordinterval :: "('i::len word \ nat) \ 'i wordinterval" where "ipcidr_tuple_to_wordinterval iprng \ iprange_interval (ipcidr_to_interval iprng)" lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval: "wordinterval_to_set (ipcidr_tuple_to_wordinterval (b, m)) = ipset_from_cidr b m" unfolding ipcidr_tuple_to_wordinterval_def ipset_from_cidr_ipcidr_to_interval ipcidr_to_interval_def by(simp add: iprange_interval.simps) lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval_uncurry: "wordinterval_to_set (ipcidr_tuple_to_wordinterval ipcidr) = uncurry ipset_from_cidr ipcidr" by(cases ipcidr, simp add: wordinterval_to_set_ipcidr_tuple_to_wordinterval) definition ipcidr_union_set :: "('i::len word \ nat) set \ ('i word) set" where "ipcidr_union_set ips \ \(base, len) \ ips. ipset_from_cidr base len" lemma ipcidr_union_set_uncurry: "ipcidr_union_set ips = (\ ipcidr \ ips. uncurry ipset_from_cidr ipcidr)" by(simp add: ipcidr_union_set_def uncurry_case_stmt) subsection\Clever Operations on IP Addresses in CIDR Notation\ text\Intersecting two intervals may result in a new interval. Example: \{1..10} \ {5..20} = {5..10}\ Intersecting two IP address ranges represented as CIDR ranges results either in the empty set or the smaller of the two ranges. It will never create a new range. \ context begin (*contributed by Lars Noschinski*) private lemma less_and_not_mask_eq: fixes i :: "('a :: len) word" assumes "r2 \ r1" "i && ~~ (mask r2) = x && ~~ (mask r2)" shows "i && ~~ (mask r1) = x && ~~ (mask r1)" proof - have "i AND NOT (mask r1) = (i && ~~ (mask r2)) && ~~ (mask r1)" (is "_ = ?w && _") using \r2 \ r1\ by (simp add: and_not_mask_twice max_def) also have "?w = x && ~~ (mask r2)" by fact also have "\ && ~~ (mask r1) = x && ~~ (mask r1)" using \r2 \ r1\ by (simp add: and_not_mask_twice max_def) finally show ?thesis . qed lemma ip_cidr_set_less: fixes i :: "'i::len word" shows "r1 \ r2 \ ip_cidr_set i r2 \ ip_cidr_set i r1" unfolding ip_cidr_set_def apply auto apply (rule less_and_not_mask_eq[where ?r2.0="LENGTH('i) - r2"]) apply auto done private lemma ip_cidr_set_intersect_subset_helper: fixes i1 r1 i2 r2 assumes disj: "ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 \ {}" and "r1 \ r2" shows "ip_cidr_set i2 r2 \ ip_cidr_set i1 r1" proof - from disj obtain j where "j \ ip_cidr_set i1 r1" "j \ ip_cidr_set i2 r2" by auto with \r1 \ r2\ have "j \ ip_cidr_set j r1" "j \ ip_cidr_set j r1" using ip_cidr_set_change_base ip_cidr_set_less by blast+ show "ip_cidr_set i2 r2 \ ip_cidr_set i1 r1" proof fix i assume "i \ ip_cidr_set i2 r2" with \j \ ip_cidr_set i2 r2\ have "i \ ip_cidr_set j r2" using ip_cidr_set_change_base by auto also have "ip_cidr_set j r2 \ ip_cidr_set j r1" using \r1 \ r2\ ip_cidr_set_less by blast also have "\ = ip_cidr_set i1 r1" using \j \ ip_cidr_set i1 r1\ ip_cidr_set_change_base by blast finally show "i \ ip_cidr_set i1 r1" . qed qed lemma ip_cidr_set_notsubset_empty_inter: "\ ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 \ \ ip_cidr_set i2 r2 \ ip_cidr_set i1 r1 \ ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 = {}" apply(cases "r1 \ r2") subgoal using ip_cidr_set_intersect_subset_helper by blast apply(cases "r2 \ r1") subgoal using ip_cidr_set_intersect_subset_helper by blast by(simp) end lemma ip_cidr_intersect: "\ ipset_from_cidr b2 m2 \ ipset_from_cidr b1 m1 \ \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2 \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2 = {}" apply(simp add: ipset_from_cidr_eq_ip_cidr_set) using ip_cidr_set_notsubset_empty_inter by blast text\Computing the intersection of two IP address ranges in CIDR notation\ fun ipcidr_conjunct :: "('i::len word \ nat) \ ('i word \ nat) \ ('i word \ nat) option" where "ipcidr_conjunct (base1, m1) (base2, m2) = ( if ipset_from_cidr base1 m1 \ ipset_from_cidr base2 m2 = {} then None else if ipset_from_cidr base1 m1 \ ipset_from_cidr base2 m2 then Some (base1, m1) else Some (base2, m2) )" text\Intersecting with an address with prefix length zero always yields a non-empty result.\ lemma ipcidr_conjunct_any: "ipcidr_conjunct a (x,0) \ None" "ipcidr_conjunct (y,0) b \ None" apply(cases a, simp add: ipset_from_cidr_0 ipset_from_cidr_not_empty) by(cases b, simp add: ipset_from_cidr_0 ipset_from_cidr_not_empty) lemma ipcidr_conjunct_correct: "(case ipcidr_conjunct (b1, m1) (b2, m2) of Some (bx, mx) \ ipset_from_cidr bx mx | None \ {}) = (ipset_from_cidr b1 m1) \ (ipset_from_cidr b2 m2)" apply(simp split: if_split_asm) using ip_cidr_intersect by fast declare ipcidr_conjunct.simps[simp del] subsection\Code Equations\ text\Executable definition using word intervals\ lemma ipcidr_conjunct_word[code_unfold]: "ipcidr_conjunct ips1 ips2 = ( if wordinterval_empty (wordinterval_intersection (ipcidr_tuple_to_wordinterval ips1) (ipcidr_tuple_to_wordinterval ips2)) then None else if wordinterval_subset (ipcidr_tuple_to_wordinterval ips1) (ipcidr_tuple_to_wordinterval ips2) then Some ips1 else Some ips2 )" apply(simp) apply(cases ips1, cases ips2, rename_tac b1 m1 b2 m2, simp) apply(auto simp add: wordinterval_to_set_ipcidr_tuple_to_wordinterval ipcidr_conjunct.simps split: if_split_asm) done (*with the code_unfold lemma before, this works!*) lemma "ipcidr_conjunct (0::32 word,0) (8,1) = Some (8, 1)" by eval export_code ipcidr_conjunct checking SML text\making element check executable\ lemma addr_in_ipset_from_netmask_code[code_unfold]: "addr \ (ipset_from_netmask base netmask) \ (base AND netmask) \ addr \ addr \ (base AND netmask) OR (NOT netmask)" by(simp add: ipset_from_netmask_def Let_def) lemma addr_in_ipset_from_cidr_code[code_unfold]: "(addr::'i::len word) \ (ipset_from_cidr pre len) \ (pre AND ((mask len) << (LENGTH('i) - len))) \ addr \ addr \ pre OR (mask (LENGTH('i) - len))" unfolding ipset_from_cidr_alt by simp end diff --git a/thys/IP_Addresses/WordInterval.thy b/thys/IP_Addresses/WordInterval.thy --- a/thys/IP_Addresses/WordInterval.thy +++ b/thys/IP_Addresses/WordInterval.thy @@ -1,775 +1,777 @@ (* Title: WordInterval.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory WordInterval -imports Main + imports + Main "Word_Lib.Word_Lemmas" + "Word_Lib.Next_and_Prev" begin section\WordInterval: Executable datatype for Machine Word Sets\ text\Stores ranges of machine words as interval. This has been proven quite efficient for IP Addresses.\ (*NOTE: All algorithms here use a straight-forward implementation. There is a lot of room for improving the computation complexity, for example by making the WordInterval a balanced, sorted tree.*) subsection\Syntax\ context notes [[typedef_overloaded]] begin datatype ('a::len) wordinterval = WordInterval "('a::len) word" \ \start (inclusive)\ "('a::len) word" \ \end (inclusive)\ | RangeUnion "'a wordinterval" "'a wordinterval" end subsection\Semantics\ fun wordinterval_to_set :: "'a::len wordinterval \ ('a::len word) set" where "wordinterval_to_set (WordInterval start end) = {start .. end}" | "wordinterval_to_set (RangeUnion r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" (*Note: The runtime of all the operations could be improved, for example by keeping the tree sorted and balanced.*) subsection\Basic operations\ text\\\\\ fun wordinterval_element :: "'a::len word \ 'a::len wordinterval \ bool" where "wordinterval_element el (WordInterval s e) \ s \ el \ el \ e" | "wordinterval_element el (RangeUnion r1 r2) \ wordinterval_element el r1 \ wordinterval_element el r2" lemma wordinterval_element_set_eq[simp]: "wordinterval_element el rg = (el \ wordinterval_to_set rg)" by(induction rg rule: wordinterval_element.induct) simp_all definition wordinterval_union :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_union r1 r2 = RangeUnion r1 r2" lemma wordinterval_union_set_eq[simp]: "wordinterval_to_set (wordinterval_union r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" unfolding wordinterval_union_def by simp fun wordinterval_empty :: "'a::len wordinterval \ bool" where "wordinterval_empty (WordInterval s e) \ e < s" | "wordinterval_empty (RangeUnion r1 r2) \ wordinterval_empty r1 \ wordinterval_empty r2" lemma wordinterval_empty_set_eq[simp]: "wordinterval_empty r \ wordinterval_to_set r = {}" by(induction r) auto definition Empty_WordInterval :: "'a::len wordinterval" where "Empty_WordInterval \ WordInterval 1 0" lemma wordinterval_empty_Empty_WordInterval: "wordinterval_empty Empty_WordInterval" by(simp add: Empty_WordInterval_def) lemma Empty_WordInterval_set_eq[simp]: "wordinterval_to_set Empty_WordInterval = {}" by(simp add: Empty_WordInterval_def) subsection\WordInterval and Lists\ text\A list of \(start, end)\ tuples.\ text\wordinterval to list\ fun wi2l :: "'a::len wordinterval \ ('a::len word \ 'a::len word) list" where "wi2l (RangeUnion r1 r2) = wi2l r1 @ wi2l r2" | "wi2l (WordInterval s e) = (if e < s then [] else [(s,e)])" text\list to wordinterval\ fun l2wi :: "('a::len word \ 'a word) list \ 'a wordinterval" where "l2wi [] = Empty_WordInterval" | "l2wi [(s,e)] = (WordInterval s e)" | "l2wi ((s,e)#rs) = (RangeUnion (WordInterval s e) (l2wi rs))" lemma l2wi_append: "wordinterval_to_set (l2wi (l1@l2)) = wordinterval_to_set (l2wi l1) \ wordinterval_to_set (l2wi l2)" proof(induction l1 arbitrary: l2 rule:l2wi.induct) case 1 thus ?case by simp next case (2 s e l2) thus ?case by (cases l2) simp_all next case 3 thus ?case by force qed lemma l2wi_wi2l[simp]: "wordinterval_to_set (l2wi (wi2l r)) = wordinterval_to_set r" by(induction r) (simp_all add: l2wi_append) lemma l2wi: "wordinterval_to_set (l2wi l) = (\ (i,j) \ set l. {i .. j})" by(induction l rule: l2wi.induct, simp_all) lemma wi2l: "(\(i,j)\set (wi2l r). {i .. j}) = wordinterval_to_set r" by(induction r rule: wi2l.induct, simp_all) lemma l2wi_remdups[simp]: "wordinterval_to_set (l2wi (remdups ls)) = wordinterval_to_set (l2wi ls)" by(simp add: l2wi) lemma wi2l_empty[simp]: "wi2l Empty_WordInterval = []" unfolding Empty_WordInterval_def by simp subsection\Optimizing and minimizing @{typ "('a::len) wordinterval"}s\ text\Removing empty intervals\ context begin fun wordinterval_optimize_empty :: "'a::len wordinterval \ 'a wordinterval" where "wordinterval_optimize_empty (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty r1o then r2o else if wordinterval_empty r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty r = r" lemma wordinterval_optimize_empty_set_eq[simp]: "wordinterval_to_set (wordinterval_optimize_empty r) = wordinterval_to_set r" by(induction r) (simp_all add: Let_def) lemma wordinterval_optimize_empty_double: "wordinterval_optimize_empty (wordinterval_optimize_empty r) = wordinterval_optimize_empty r" by(induction r) (simp_all add: Let_def) private fun wordinterval_empty_shallow :: "'a::len wordinterval \ bool" where "wordinterval_empty_shallow (WordInterval s e) \ e < s" | "wordinterval_empty_shallow (RangeUnion _ _) \ False" private lemma helper_optimize_shallow: "wordinterval_empty_shallow (wordinterval_optimize_empty r) = wordinterval_empty (wordinterval_optimize_empty r)" by(induction r) fastforce+ private fun wordinterval_optimize_empty2 where "wordinterval_optimize_empty2 (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty_shallow r1o then r2o else if wordinterval_empty_shallow r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty2 r = r" lemma wordinterval_optimize_empty_code[code_unfold]: "wordinterval_optimize_empty = wordinterval_optimize_empty2" by (subst fun_eq_iff, clarify, rename_tac r, induct_tac r) (unfold wordinterval_optimize_empty.simps wordinterval_optimize_empty2.simps Let_def helper_optimize_shallow, simp_all) end text\Merging overlapping intervals\ context begin private definition disjoint :: "'a set \ 'a set \ bool" where "disjoint A B \ A \ B = {}" private primrec interval_of :: "('a::len) word \ 'a word \ 'a word set" where "interval_of (s,e) = {s .. e}" declare interval_of.simps[simp del] private definition disjoint_intervals :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "disjoint_intervals A B \ disjoint (interval_of A) (interval_of B)" private definition not_disjoint_intervals :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "not_disjoint_intervals A B \ \ disjoint (interval_of A) (interval_of B)" private lemma [code]: "not_disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s \ e' \ s' \ e \ s \ e \ s' \ e')" apply(cases A, cases B) apply(simp add: not_disjoint_intervals_def interval_of.simps disjoint_def) done private lemma [code]: "disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s > e' \ s' > e \ s > e \ s' > e')" apply(cases A, cases B) apply(simp add: disjoint_intervals_def interval_of.simps disjoint_def) by fastforce text\BEGIN merging overlapping intervals\ (*result has no empty intervals and all are disjoint. merging things such as [1,7] [8,10] would still be possible*) private fun merge_overlap :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_overlap s [] = [s]" | "merge_overlap (s,e) ((s',e')#ss) = ( if not_disjoint_intervals (s,e) (s',e') then (min s s', max e e')#ss else (s',e')#merge_overlap (s,e) ss)" private lemma not_disjoint_union: fixes s :: "('a::len) word" shows "\ disjoint {s..e} {s'..e'} \ {s..e} \ {s'..e'} = {min s s' .. max e e'}" by(auto simp add: disjoint_def min_def max_def) private lemma disjoint_subset: "disjoint A B \ A \ B \ C \ A \ C" unfolding disjoint_def by blast private lemma merge_overlap_helper1: "interval_of A \ (\s \ set ss. interval_of s) \ (\s \ set (merge_overlap A ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(drule_tac C="(\x\set xs. interval_of x)" in disjoint_subset) apply(simp_all) done private lemma merge_overlap_helper2: "\s'\set ss. \ disjoint (interval_of A) (interval_of s') \ interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_overlap A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(simp) by blast private lemma merge_overlap_length: "\s' \ set ss. \ disjoint (interval_of A) (interval_of s') \ length (merge_overlap A ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) done lemma "merge_overlap (1:: 16 word,2) [(1, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(2, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(3, 7)] = [(3, 7), (1,2)]" by eval private function listwordinterval_compress :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_compress [] = []" | "listwordinterval_compress (s#ss) = ( if \s' \ set ss. disjoint_intervals s s' then s#listwordinterval_compress ss else listwordinterval_compress (merge_overlap s ss))" by(pat_completeness, auto) private termination listwordinterval_compress apply (relation "measure length") apply(rule wf_measure) apply(simp) using disjoint_intervals_def merge_overlap_length by fastforce private lemma listwordinterval_compress: "(\s \ set (listwordinterval_compress ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_compress.induct) apply(simp) apply(simp) apply(intro impI) apply(simp add: disjoint_intervals_def) apply(drule merge_overlap_helper2) apply(simp) done lemma "listwordinterval_compress [(1::32 word,3), (8,10), (2,5), (3,7)] = [(8, 10), (1, 7)]" by eval private lemma A_in_listwordinterval_compress: "A \ set (listwordinterval_compress ss) \ interval_of A \ (\s \ set ss. interval_of s)" using listwordinterval_compress by blast private lemma listwordinterval_compress_disjoint: "A \ set (listwordinterval_compress ss) \ B \ set (listwordinterval_compress ss) \ A \ B \ disjoint (interval_of A) (interval_of B)" apply(induction ss arbitrary: rule: listwordinterval_compress.induct) apply(simp) apply(simp split: if_split_asm) apply(elim disjE) apply(simp_all) apply(simp_all add: disjoint_intervals_def disjoint_def) apply(blast dest: A_in_listwordinterval_compress)+ done text\END merging overlapping intervals\ text\BEGIN merging adjacent intervals\ private fun merge_adjacent :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_adjacent s [] = [s]" | "merge_adjacent (s,e) ((s',e')#ss) = ( if s \e \ s' \ e' \ word_next e = s' then (s, e')#ss else if s \e \ s' \ e' \ word_next e' = s then (s', e)#ss else (s',e')#merge_adjacent (s,e) ss)" private lemma merge_adjacent_helper: "interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_adjacent A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: interval_of.simps) apply(intro impI conjI) apply (metis Un_assoc word_adjacent_union) apply(elim conjE) apply(drule(2) word_adjacent_union) subgoal by (blast) subgoal by (metis word_adjacent_union Un_assoc) by blast private lemma merge_adjacent_length: "\(s', e')\set ss. s \ e \ s' \ e' \ (word_next e = s' \ word_next e' = s) \ length (merge_adjacent (s,e) ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(case_tac x) apply(simp add: ) by blast private function listwordinterval_adjacent :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_adjacent [] = []" | "listwordinterval_adjacent ((s,e)#ss) = ( if \(s',e') \ set ss. \ (s \e \ s' \ e' \ (word_next e = s' \ word_next e' = s)) then (s,e)#listwordinterval_adjacent ss else listwordinterval_adjacent (merge_adjacent (s,e) ss))" by(pat_completeness, auto) private termination listwordinterval_adjacent apply (relation "measure length") apply(rule wf_measure) apply(simp) apply(simp) using merge_adjacent_length by fastforce private lemma listwordinterval_adjacent: "(\s \ set (listwordinterval_adjacent ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_adjacent.induct) apply(simp) apply(simp add: merge_adjacent_helper) done lemma "listwordinterval_adjacent [(1::16 word, 3), (5, 10), (10,10), (4,4)] = [(10, 10), (1, 10)]" by eval text\END merging adjacent intervals\ definition wordinterval_compress :: "('a::len) wordinterval \ 'a wordinterval" where "wordinterval_compress r \ l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))" text\Correctness: Compression preserves semantics\ lemma wordinterval_compress: "wordinterval_to_set (wordinterval_compress r) = wordinterval_to_set r" unfolding wordinterval_compress_def proof - have interval_of': "interval_of s = (case s of (s,e) \ {s .. e})" for s by (cases s) (simp add: interval_of.simps) have "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = (\x\set (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))). interval_of x)" by (force simp: interval_of' l2wi) also have "\ = (\s\set (wi2l (wordinterval_optimize_empty r)). interval_of s)" by(simp add: listwordinterval_compress listwordinterval_adjacent) also have "\ = (\(i, j)\set (wi2l (wordinterval_optimize_empty r)). {i..j})" by(simp add: interval_of') also have "\ = wordinterval_to_set r" by(simp add: wi2l) finally show "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = wordinterval_to_set r" . qed end text\Example\ lemma "(wi2l \ (wordinterval_compress :: 32 wordinterval \ 32 wordinterval) \ l2wi) [(70, 80001), (0,0), (150, 8000), (1,3), (42,41), (3,7), (56, 200), (8,10)] = [(56, 80001), (0, 10)]" by eval lemma "wordinterval_compress (RangeUnion (RangeUnion (WordInterval (1::32 word) 5) (WordInterval 8 10)) (WordInterval 3 7)) = WordInterval 1 10" by eval subsection\Further operations\ text\\\\\ definition wordinterval_Union :: "('a::len) wordinterval list \ 'a wordinterval" where "wordinterval_Union ws = wordinterval_compress (foldr wordinterval_union ws Empty_WordInterval)" lemma wordinterval_Union: "wordinterval_to_set (wordinterval_Union ws) = (\ w \ (set ws). wordinterval_to_set w)" by(induction ws) (simp_all add: wordinterval_compress wordinterval_Union_def) context begin private fun wordinterval_setminus' :: "'a::len wordinterval \ 'a wordinterval \ 'a wordinterval" where "wordinterval_setminus' (WordInterval s e) (WordInterval ms me) = ( if s > e \ ms > me then WordInterval s e else if me \ e then WordInterval (if ms = 0 then 1 else s) (min e (word_prev ms)) else if ms \ s then WordInterval (max s (word_next me)) (if me = max_word then 0 else e) else RangeUnion (WordInterval (if ms = 0 then 1 else s) (word_prev ms)) (WordInterval (word_next me) (if me = max_word then 0 else e)) )" | "wordinterval_setminus' (RangeUnion r1 r2) t = RangeUnion (wordinterval_setminus' r1 t) (wordinterval_setminus' r2 t)"| "wordinterval_setminus' t (RangeUnion r1 r2) = wordinterval_setminus' (wordinterval_setminus' t r1) r2" private lemma wordinterval_setminus'_rr_set_eq: "wordinterval_to_set(wordinterval_setminus' (WordInterval s e) (WordInterval ms me)) = wordinterval_to_set (WordInterval s e) - wordinterval_to_set (WordInterval ms me)" apply(simp only: wordinterval_setminus'.simps) apply(case_tac "e < s") apply simp apply(case_tac "me < ms") apply simp apply(case_tac [!] "e \ me") apply(case_tac [!] "ms = 0") apply(case_tac [!] "ms \ s") apply(case_tac [!] "me = max_word") apply(simp_all add: word_next_unfold word_prev_unfold min_def max_def) apply(safe) apply(auto) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) done private lemma wordinterval_setminus'_set_eq: "wordinterval_to_set (wordinterval_setminus' r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" apply(induction rule: wordinterval_setminus'.induct) using wordinterval_setminus'_rr_set_eq apply blast apply auto done lemma wordinterval_setminus'_empty_struct: "wordinterval_empty r2 \ wordinterval_setminus' r1 r2 = r1" by(induction r1 r2 rule: wordinterval_setminus'.induct) auto definition wordinterval_setminus :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_setminus r1 r2 = wordinterval_compress (wordinterval_setminus' r1 r2)" lemma wordinterval_setminus_set_eq[simp]: "wordinterval_to_set (wordinterval_setminus r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" by(simp add: wordinterval_setminus_def wordinterval_compress wordinterval_setminus'_set_eq) end definition wordinterval_UNIV :: "'a::len wordinterval" where "wordinterval_UNIV \ WordInterval 0 max_word" lemma wordinterval_UNIV_set_eq[simp]: "wordinterval_to_set wordinterval_UNIV = UNIV" unfolding wordinterval_UNIV_def using max_word_max by fastforce fun wordinterval_invert :: "'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_invert r = wordinterval_setminus wordinterval_UNIV r" lemma wordinterval_invert_set_eq[simp]: "wordinterval_to_set (wordinterval_invert r) = UNIV - wordinterval_to_set r" by(auto) lemma wordinterval_invert_UNIV_empty: "wordinterval_empty (wordinterval_invert wordinterval_UNIV)" by simp lemma wi2l_univ[simp]: "wi2l wordinterval_UNIV = [(0, max_word)]" unfolding wordinterval_UNIV_def by simp text\\\\\ context begin private lemma "{(s::nat) .. e} \ {s' .. e'} = {} \ s > e' \ s' > e \ s > e \ s' > e'" by simp linarith private fun wordinterval_intersection' :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection' (WordInterval s e) (WordInterval s' e') = ( if s > e \ s' > e' \ s > e' \ s' > e \ s > e \ s' > e' then Empty_WordInterval else WordInterval (max s s') (min e e') )" | "wordinterval_intersection' (RangeUnion r1 r2) t = RangeUnion (wordinterval_intersection' r1 t) (wordinterval_intersection' r2 t)"| "wordinterval_intersection' t (RangeUnion r1 r2) = RangeUnion (wordinterval_intersection' t r1) (wordinterval_intersection' t r2)" private lemma wordinterval_intersection'_set_eq: "wordinterval_to_set (wordinterval_intersection' r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(induction r1 r2 rule: wordinterval_intersection'.induct) (auto) lemma "wordinterval_intersection' (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = RangeUnion (RangeUnion (WordInterval 1 3) (WordInterval 1 0)) (WordInterval 1 3)" by eval definition wordinterval_intersection :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection r1 r2 \ wordinterval_compress (wordinterval_intersection' r1 r2)" lemma wordinterval_intersection_set_eq[simp]: "wordinterval_to_set (wordinterval_intersection r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(simp add: wordinterval_intersection_def wordinterval_compress wordinterval_intersection'_set_eq) lemma "wordinterval_intersection (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = WordInterval 1 3" by eval end definition wordinterval_subset :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_subset r1 r2 \ wordinterval_empty (wordinterval_setminus r1 r2)" lemma wordinterval_subset_set_eq[simp]: "wordinterval_subset r1 r2 = (wordinterval_to_set r1 \ wordinterval_to_set r2)" unfolding wordinterval_subset_def by simp definition wordinterval_eq :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_eq r1 r2 = (wordinterval_subset r1 r2 \ wordinterval_subset r2 r1)" lemma wordinterval_eq_set_eq: "wordinterval_eq r1 r2 \ wordinterval_to_set r1 = wordinterval_to_set r2" unfolding wordinterval_eq_def by auto thm iffD1[OF wordinterval_eq_set_eq] (*declare iffD1[OF wordinterval_eq_set_eq, simp]*) lemma wordinterval_eq_comm: "wordinterval_eq r1 r2 \ wordinterval_eq r2 r1" unfolding wordinterval_eq_def by fast lemma wordinterval_to_set_alt: "wordinterval_to_set r = {x. wordinterval_element x r}" unfolding wordinterval_element_set_eq by blast lemma wordinterval_un_empty: "wordinterval_empty r1 \ wordinterval_eq (wordinterval_union r1 r2) r2" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_un_emty_b: "wordinterval_empty r2 \ wordinterval_eq (wordinterval_union r1 r2) r1" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_Diff_triv: "wordinterval_empty (wordinterval_intersection a b) \ wordinterval_eq (wordinterval_setminus a b) a" unfolding wordinterval_eq_set_eq by simp blast text\A size of the datatype, does not correspond to the cardinality of the corresponding set\ fun wordinterval_size :: "('a::len) wordinterval \ nat" where "wordinterval_size (RangeUnion a b) = wordinterval_size a + wordinterval_size b" | "wordinterval_size (WordInterval s e) = (if s \ e then 1 else 0)" lemma wordinterval_size_length: "wordinterval_size r = length (wi2l r)" by(induction r) (auto) lemma Ex_wordinterval_nonempty: "\x::('a::len wordinterval). y \ wordinterval_to_set x" proof show "y \ wordinterval_to_set wordinterval_UNIV" by simp qed lemma wordinterval_eq_reflp: "reflp wordinterval_eq" apply(rule reflpI) by(simp only: wordinterval_eq_set_eq) lemma wordintervalt_eq_symp: "symp wordinterval_eq" apply(rule sympI) by(simp add: wordinterval_eq_comm) lemma wordinterval_eq_transp: "transp wordinterval_eq" apply(rule transpI) by(simp only: wordinterval_eq_set_eq) lemma wordinterval_eq_equivp: "equivp wordinterval_eq" by (auto intro: equivpI wordinterval_eq_reflp wordintervalt_eq_symp wordinterval_eq_transp) text\The smallest element in the interval\ definition is_lowest_element :: "'a::ord \ 'a set \ bool" where "is_lowest_element x S = (x \ S \ (\y\S. y \ x \ y = x))" lemma fixes x :: "'a :: complete_lattice" assumes "x \ S" shows " x = Inf S \ is_lowest_element x S" using assms apply(simp add: is_lowest_element_def) by (simp add: Inf_lower eq_iff) lemma fixes x :: "'a :: linorder" assumes "finite S" and "x \ S" shows "is_lowest_element x S \ x = Min S" apply(rule) subgoal apply(simp add: is_lowest_element_def) apply(subst Min_eqI[symmetric]) using assms by(auto) by (metis Min.coboundedI assms(1) assms(2) dual_order.antisym is_lowest_element_def) text\Smallest element in the interval\ fun wordinterval_lowest_element :: "'a::len wordinterval \ 'a word option" where "wordinterval_lowest_element (WordInterval s e) = (if s \ e then Some s else None)" | "wordinterval_lowest_element (RangeUnion A B) = (case (wordinterval_lowest_element A, wordinterval_lowest_element B) of (Some a, Some b) \ Some (if a < b then a else b) | (None, Some b) \ Some b | (Some a, None) \ Some a | (None, None) \ None)" lemma wordinterval_lowest_none_empty: "wordinterval_lowest_element r = None \ wordinterval_empty r" proof(induction r) case WordInterval thus ?case by simp next case RangeUnion thus ?case by fastforce qed lemma wordinterval_lowest_element_correct_A: "wordinterval_lowest_element r = Some x \ is_lowest_element x (wordinterval_to_set r)" unfolding is_lowest_element_def apply(induction r arbitrary: x rule: wordinterval_lowest_element.induct) apply(rename_tac rs re x, case_tac "rs \ re", auto)[1] apply(subst(asm) wordinterval_lowest_element.simps(2)) apply(rename_tac A B x) apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(simp_all add: wordinterval_lowest_none_empty)[3] apply fastforce done lemma wordinterval_lowest_element_set_eq: assumes "\ wordinterval_empty r" shows "(wordinterval_lowest_element r = Some x) = (is_lowest_element x (wordinterval_to_set r))" (*unfolding is_lowest_element_def*) proof(rule iffI) assume "wordinterval_lowest_element r = Some x" thus "is_lowest_element x (wordinterval_to_set r)" using wordinterval_lowest_element_correct_A wordinterval_lowest_none_empty by simp next assume "is_lowest_element x (wordinterval_to_set r)" with assms show "(wordinterval_lowest_element r = Some x)" proof(induction r arbitrary: x rule: wordinterval_lowest_element.induct) case 1 thus ?case by(simp add: is_lowest_element_def) next case (2 A B x) have is_lowest_RangeUnion: "is_lowest_element x (wordinterval_to_set A \ wordinterval_to_set B) \ is_lowest_element x (wordinterval_to_set A) \ is_lowest_element x (wordinterval_to_set B)" by(simp add: is_lowest_element_def) (*why \ A B?*) have wordinterval_lowest_element_RangeUnion: "\a b A B. wordinterval_lowest_element A = Some a \ wordinterval_lowest_element B = Some b \ wordinterval_lowest_element (RangeUnion A B) = Some (min a b)" by(auto dest!: wordinterval_lowest_element_correct_A simp add: is_lowest_element_def min_def) from 2 show ?case apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(auto simp add: is_lowest_element_def)[3] apply(subgoal_tac "\ wordinterval_empty A \ \ wordinterval_empty B") prefer 2 using arg_cong[where f = Not, OF wordinterval_lowest_none_empty] apply blast apply(drule(1) wordinterval_lowest_element_RangeUnion) apply(simp split: option.split_asm add: min_def) apply(drule is_lowest_RangeUnion) apply(elim disjE) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) using wordinterval_lowest_element_correct_A[simplified is_lowest_element_def] by (metis Un_iff not_le) qed qed text\Cardinality approximation for @{typ "('a::len) wordinterval"}s\ context begin lemma card_atLeastAtMost_word: fixes s::"('a::len) word" shows "card {s..e} = Suc (unat e) - (unat s)" apply(cases "s > e") apply(simp) apply(subst(asm) Word.word_less_nat_alt) apply simp apply(subst upto_enum_set_conv2[symmetric]) apply(subst List.card_set) apply(simp add: remdups_enum_upto) done fun wordinterval_card :: "('a::len) wordinterval \ nat" where "wordinterval_card (WordInterval s e) = Suc (unat e) - (unat s)" | "wordinterval_card (RangeUnion a b) = wordinterval_card a + wordinterval_card b" lemma wordinterval_card: "wordinterval_card r \ card (wordinterval_to_set r)" proof(induction r) case WordInterval thus ?case by (simp add: card_atLeastAtMost_word) next case (RangeUnion r1 r2) have "card (wordinterval_to_set r1 \ wordinterval_to_set r2) \ card (wordinterval_to_set r1) + card (wordinterval_to_set r2)" using Finite_Set.card_Un_le by blast with RangeUnion show ?case by(simp) qed text\With @{thm wordinterval_compress} it should be possible to get the exact cardinality\ end end diff --git a/thys/Interval_Arithmetic_Word32/Interpreter.thy b/thys/Interval_Arithmetic_Word32/Interpreter.thy --- a/thys/Interval_Arithmetic_Word32/Interpreter.thy +++ b/thys/Interval_Arithmetic_Word32/Interpreter.thy @@ -1,601 +1,598 @@ text\ Interpreter.thy defines a simple programming language over interval-valued variables and executable semantics (interpreter) for that language. We then prove that the interpretation of interval terms is a sound over-approximation of a real-valued semantics of the same language. Our language is a version of first order dynamic logic-style regular programs. We use a finite identifier space for compatibility with Differential-Dynamic-Logic, where identifier finiteness is required to treat program states as Banach spaces to enable differentiation. \ (* Author: Brandon Bohrer *) theory Interpreter imports Complex_Main Finite_String Interval_Word32 - Word_Lib.Word_Lemmas - Word_Lib.Word_Lib - Word_Lib.Word_Syntax begin section\Syntax\ text\Our term language supports variables, polynomial arithmetic, and extrema. This choice was made based on the needs of the original paper and could be extended if necessary.\ datatype trm = Var fin_string | Const lit | Plus trm trm | Times trm trm | Neg trm | Max trm trm | Min trm trm | Abs trm text\Our statement language is nondeterministic first-order regular programs. This coincides with the discrete subset of hybrid programs from the dL entry. Our assertion language are the formulas of first-order dynamic logic\ datatype prog = Assign fin_string "trm" (infixr ":=" 10) | AssignAny fin_string | Test "formula" ("?") | Choice "prog" "prog" (infixl "\\" 10) | Sequence "prog" "prog" (infixr ";;" 8) | Loop "prog" ("_**") and formula = Geq "trm" "trm" | Not "formula" ("!") | And "formula" "formula" (infixl "&&" 8) | Exists fin_string "formula" | Diamond "prog" "formula" ("(\ _ \ _)" 10) text\Derived forms\ definition Or :: "formula \ formula \ formula" (infixl "||" 7) where or_simp[simp]:"Or P Q = Not (And (Not P) (Not Q))" definition Equals :: "trm \ trm \ formula" where equals_simp[simp]:"Equals \ \' = (And (Geq \ \') (Geq \' \))" definition Greater :: "trm \ trm \ formula" where greater_simp[simp]:"Greater \ \' = Not (Geq \' \)" definition Leq :: "trm \ trm \ formula" where leq_simp[simp]:"Leq \ \' = (Geq \' \)" definition Less :: "trm \ trm \ formula" where less_simp[simp]:"Less \ \' = (Not (Geq \ \'))" section \Semantics\ text\ States over reals vs. word intervals which contain them\ type_synonym rstate = "fin_string \ real" type_synonym wstate = "(fin_string + fin_string) \ word" definition wstate::"wstate \ prop" where wstate_def[simp]:"wstate \ \ (\i. word (\ (Inl i)) \ word (\ (Inr i)))" text\Interpretation of a term in a state\ inductive rtsem :: "trm \ rstate \ real \ bool" ("([_]_ \ _)" 10) where rtsem_Const:"Rep_bword w \\<^sub>E r \ ([Const w]\ \ r)" | rtsem_Var:"([Var x]\ \ \ x)" | rtsem_Plus:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ ([Plus \\<^sub>1 \\<^sub>2]\ \ (r1 + r2))" | rtsem_Times:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ ([Times \\<^sub>1 \\<^sub>2]\ \ (r1 * r2))" | rtsem_Max:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ ([Max \\<^sub>1 \\<^sub>2]\ \ (max r1 r2))" | rtsem_Min:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ ([Min \\<^sub>1 \\<^sub>2]\ \ (min r1 r2))" | rtsem_Abs:"\([\\<^sub>1]\ \ r1)\ \ ([Abs \\<^sub>1]\ \ (abs r1))" | rtsem_Neg:"([\]\ \ r) \ ([Neg \]\ \ -r)" inductive_simps rtsem_Const_simps[simp] : "([(Const w)]\ \ r)" and rtsem_Var_simps[simp] : "([Var x]\ \ r)" and rtsem_PlusU_simps[simp] : "([Plus \\<^sub>1 \\<^sub>2]\ \ r)" and rtsem_TimesU_simps[simp] : "([Times \\<^sub>1 \\<^sub>2]\ \ r)" and rtsem_Max_simps[simp] : "([Max \\<^sub>1 \\<^sub>2] \ \ r)" and rtsem_Min_simps[simp] : "([Min \\<^sub>1 \\<^sub>2] \ \ r)" and rtsem_Abs_simps[simp] : "([Abs \] \ \ r)" and rtsem_Neg_simps[simp] : "([Neg \] \ \ r)" definition set_less :: "real set \ real set \ bool" (infix "<\<^sub>S" 10) where "set_less A B \ (\ x y. x \ A \ y \ B \ x < y)" definition set_geq :: "real set \ real set \ bool" (infix "\\<^sub>S" 10) where "set_geq A B \ (\ x y. x \ A \ y \ B \ x \ y)" text\Interpretation of an assertion in a state\ inductive rfsem :: "formula \ rstate \ bool \ bool" ("([_]_) \ _" 20) where rGreaterT:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ r1 > r2 \ ([Greater \\<^sub>1 \\<^sub>2] \ \ True)" | rGreaterF:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ r2 \ r1 \ ([Greater \\<^sub>1 \\<^sub>2] \ \ False)" | rGeqT:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ r1 \ r2 \ ([Geq \\<^sub>1 \\<^sub>2] \ \ True)" | rGeqF:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ r2 > r1 \ ([Geq \\<^sub>1 \\<^sub>2] \ \ False)" | rEqualsT:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ r1 = r2 \ ([Equals \\<^sub>1 \\<^sub>2] \ \ True)" | rEqualsF:"\([\\<^sub>1]\ \ r1); ([\\<^sub>2]\ \ r2)\ \ r1 \ r2 \ ([Equals \\<^sub>1 \\<^sub>2] \ \ False)" | rAndT:"\([\]\ \ True); ([\]\ \ True)\ \ ([And \ \]\ \ True)" | rAndF1:"([\]\ \ False) \ ([And \ \]\ \ False)" | rAndF2:"([\]\ \ False) \ ([And \ \]\ \ False)" | rOrT1:"([\]\ \ True) \ ([Or \ \]\ \ True)" | rOrT2:"([\]\ \ True) \ ([Or \ \]\ \ True)" | rOrF:"\([\]\ \ False) ;([\]\ \ False)\ \ ([And \ \]\ \ False)" | rNotT:"([\]\ \ False) \ ([(Not \)]\ \ True)" | rNotF:"([\]\ \ True) \ ([(Not \)]\ \ False)" inductive_simps rfsem_Greater_simps[simp]: "([Greater \\<^sub>1 \\<^sub>2]\ \ b)" and rfsem_Geq_simps[simp]: "([Geq \\<^sub>1 \\<^sub>2]\ \ b)" and rfsem_Equals_simps[simp]: "([Equals \\<^sub>1 \\<^sub>2]\ \ b)" and rfsem_And_simps[simp]: "([And \ \]\ \ b)" and rfsem_Or_simps[simp]: "([(Or \ \)]\ \ b)" and rfsem_Not_simps[simp]: "([Not \]\ \ b)" text\Interpretation of a program is a transition relation on states\ inductive rpsem :: "prog \ rstate \ rstate \ bool" ("([_]_) \ _" 20) where rTest[simp]:"\([\]\ \ True); \ = \\ \ ([? \]\ \ \)" | rSeq[simp]:"\([\]\ \ \); ([\]\ \ \)\ \ ([\;; \]\ \ \)" | rAssign[simp]:"\([\]\ \ r); \ = (\ (x := r))\ \ ([Assign x \]\ \ \)" | rChoice1[simp]:"([\]\ \ \) \ ([Choice \ \]\ \ \)" | rChoice2[simp]:"([\]\ \ \) \ ([Choice \ \]\ \ \)" inductive_simps rpsem_Test_simps[simp]: "([? \]\ \ b)" and rpsem_Seq_simps[simp]: "([\;; \]\ \ b)" and rpsem_Assign_simps[simp]: "([Assign x \]\ \ b)" and rpsem_Choice_simps[simp]: "([Choice \ \]\ \ b)" text\Upper bound of arbitrary term\ fun wtsemU :: "trm \ wstate \ word * word " ("([_]<>_)" 20) where "([Const r]<>\) = (Rep_bword r::word, Rep_bword r)" | wVarU:"([Var x]<>\) = (\ (Inl x), \ (Inr x))" | wPlusU:"([Plus \\<^sub>1 \\<^sub>2]<> \) = (let (l1, u1) = [\\<^sub>1]<> \ in let (l2, u2) = [\\<^sub>2]<> \ in (pl l1 l2, pu u1 u2))" | wTimesU:"([(Times \\<^sub>1 \\<^sub>2)]<> \) = (let (l1, u1) = [\\<^sub>1]<> \ in let (l2, u2) = [\\<^sub>2]<> \ in (tl l1 u1 l2 u2, tu l1 u1 l2 u2))" | wMaxU:"([(Max \\<^sub>1 \\<^sub>2)]<> \) = (let (l1, u1) = [\\<^sub>1]<> \ in let (l2, u2) = [\\<^sub>2]<> \ in (wmax l1 l2, wmax u1 u2))" | wMinU:"([(Min \\<^sub>1 \\<^sub>2)]<> \) = (let (l1, u1) = [\\<^sub>1]<> \ in let (l2, u2) = [\\<^sub>2]<> \ in (wmin l1 l2, wmin u1 u2))" | wNegU:"([(Neg \)]<> \) = (let (l, u) = [\]<> \ in (wneg u, wneg l))" | wAbsU:"([(Abs \\<^sub>1)]<> \) = (let (l1, u1) = [\\<^sub>1]<> \ in (wmax l1 (wneg u1), wmax u1 (wneg l1)))" inductive wfsem :: "formula \ wstate \ bool \ bool" ("([[_]]_ \ _)" 20) where wGreaterT:"wgreater (fst ([\\<^sub>1]<>\)) (snd ([\\<^sub>2]<>\)) \ ([[(Greater \\<^sub>1 \\<^sub>2)]]\ \ True)" | wGreaterF:"wgeq (fst ([\\<^sub>2]<>\)) (snd ([\\<^sub>1]<>\)) \ ([[(Greater \\<^sub>1 \\<^sub>2)]]\ \ False)" | wGeqT:"wgeq (fst ([\\<^sub>1]<> \)) (snd ([\\<^sub>2]<>\)) \ ([[(Geq \\<^sub>1 \\<^sub>2)]]\ \ True)" | wGeqF:"wgreater (fst ([\\<^sub>2]<>\)) (snd ([\\<^sub>1]<>\)) \ ([[(Geq \\<^sub>1 \\<^sub>2)]]\ \ False)" | wEqualsT:"\(fst ([\\<^sub>2]<>\) = snd ([\\<^sub>2]<>\)); (snd ([\\<^sub>2]<>\) = snd ([\\<^sub>1]<>\)); (snd ([\\<^sub>1]<>\) = fst ([\\<^sub>1]<>\)); (fst ([\\<^sub>2]<>\) \ NEG_INF); (fst ([\\<^sub>2]<>\) \ POS_INF)\ \ ([[Equals \\<^sub>1 \\<^sub>2]] \ \ True)" | wEqualsF1:"wgreater (fst ([\\<^sub>1]<> \)) (snd ([\\<^sub>2]<>\)) \ ([[Equals \\<^sub>1 \\<^sub>2]] \ \ False)" | wEqualsF2:"wgreater (fst ([\\<^sub>2]<> \)) (snd ([\\<^sub>1]<>\)) \ ([[Equals \\<^sub>1 \\<^sub>2]] \ \ False)" | wAndT:"\[[\]]\ \ True; [[\]]\ \ True\ \ ([[And \ \]]\ \ True)" | wAndF1:"[[\]]\ \ False \ ([[And \ \]]\ \ False)" | wAndF2:"[[\]]\ \ False \ ([[And \ \]]\ \ False)" | wOrT1:"([[\]]\ \ True) \ ([[Or \ \]]\ \ True)" | wOrT2:"([[\]]\ \ True) \ ([[Or \ \]]\ \ True)" | wOrF:"\[[\]]\ \ False; [[\]]\ \ False\ \ ([[And \ \]]\ \ False)" | wNotT:"([[\]]\ \ False) \ ([[Not \]]\ \ True)" | wNotF:"([[\]]\ \ True) \ ([[Not \]]\ \ False)" inductive_simps wfsem_Gr_simps[simp]: "([[Le \\<^sub>1 \\<^sub>2]]\ \ b)" and wfsem_And_simps[simp]: "([[And \ \]]\ \ b)" and wfsem_Or_simps[simp]: "([[Or \ \]]\ \ b)" and wfsem_Not_simps[simp]: "([[Not \]]\ \ b)" and wfsem_Equals_simps[simp]: "([[Equals \\<^sub>1 \\<^sub>2]]\ \ b)" text\Program semantics\ inductive wpsem :: "prog \ wstate \ wstate \ bool" ("([[_]]_ \ _)" 20) where wTest:"([[\]]\ \ True) \ \ = \ \ ([[? \]] \ \ \)" | wSeq:"([[\]]\ \ \) \ ([[\]] \ \ \) \ ([[\;; \]] \ \ \)" | wAssign:"\ = ((\ ((Inr x) := snd([\]<>\))) ((Inl x) := fst([\]<>\))) \ ([[Assign x \]] \ \ \)" | wChoice1[simp]:"([[\]]\ \ \) \ ([[Choice \ \]]\ \ \)" | wChoice2[simp]:"([[\]]\ \ \) \ ([[Choice \ \]]\ \ \)" inductive_simps wpsem_Test_simps[simp]: "([[Test \]]\ \ b)" and wpsem_Seq_simps[simp]: "([[\;; \]]\ \ b)" and wpsem_Assign_simps[simp]: "([[Assign x \]]\ \ b)" and wpsem_Choice_simps[simp]: "([[Choice \ \]]\ \ b)" lemmas real_max_mono = Lattices.linorder_class.max.mono lemmas real_minus_le_minus = Groups.ordered_ab_group_add_class.neg_le_iff_le text\Interval state consists of upper and lower bounds for each real variable\ inductive represents_state::"wstate \ rstate \ bool" (infix "REP" 10) where REPI:"(\x. (\ (Inl x) \\<^sub>L \' x) \ (\ (Inr x) \\<^sub>U \' x)) \ (\ REP \')" inductive_simps repstate_simps:"\ REP \'" section\Soundness proofs\ text\Interval term valuation soundly contains real valuation\ lemma trm_sound: fixes \::"trm" shows "([\]\' \ r) \ (\ REP \') \ ([\]<>\) \\<^sub>P r" proof (induction rule: rtsem.induct) case rtsem_Const fix w r \' show "Rep_bword w \\<^sub>E r \ \ REP \' \ [Const w]<>\ \\<^sub>P r" using repU_def repL_def repP_def repe.simps rep_simps repstate_simps by auto next case rtsem_Var fix x \' show "\ REP \' \ [Var x]<>\ \\<^sub>P \' x" by(auto simp add: repU_def repL_def repP_def repe.simps rep_simps repstate_simps) next case rtsem_Plus fix \\<^sub>1 :: "trm" and \':: "rstate" and r1 and \\<^sub>2 :: "trm" and r2 assume rep:"\ REP \'" assume eval1:"[\\<^sub>1]\' \ r1" assume "(\ REP \' \ [\\<^sub>1]<>\ \\<^sub>P r1)" then have IH1:"[\\<^sub>1]<>\ \\<^sub>P r1" using rep by auto assume eval2:"[\\<^sub>2]\' \ r2" assume "(\ REP \' \ [\\<^sub>2]<>\ \\<^sub>P r2)" then have IH2:"[\\<^sub>2]<>\ \\<^sub>P r2" using rep by auto obtain l1 u1 l2 u2 where lu1:"(l1, u1) = ([\\<^sub>1]<> \)" and lu2:"(l2, u2) = ([\\<^sub>2]<> \)" using IH1 IH2 repP_def by auto from lu1 and lu2 have lu1':"([\\<^sub>1]<> \) = (l1, u1)" and lu2':"([\\<^sub>2]<> \) = (l2, u2)" by auto have l1:"l1 \\<^sub>L r1" using IH1 lu1 unfolding repP_def by auto have u1:"u1 \\<^sub>U r1" using IH1 lu1 unfolding repP_def by auto have l2:"l2 \\<^sub>L r2" using IH2 lu2 unfolding repP_def by auto have u2:"u2 \\<^sub>U r2" using IH2 lu2 unfolding repP_def by auto then have "([(Plus \\<^sub>1 \\<^sub>2)]<>\) = (pl l1 l2, pu u1 u2)" using lu1' lu2' by auto have lBound:"(pl l1 l2 \\<^sub>L r1 + r2)" using l1 l2 pl_lemma by auto have uBound:"(pu u1 u2 \\<^sub>U r1 + r2)" using pu_lemma[OF u1 u2] by auto have "(pl l1 l2, pu u1 u2) \\<^sub>P (r1 + r2)" unfolding repP_def Let_def using lBound uBound by auto then show"[Plus \\<^sub>1 \\<^sub>2]<>\ \\<^sub>P r1 + r2" using lu1' lu2' by auto next case rtsem_Times fix \\<^sub>1 :: "trm" and \' r1 and \\<^sub>2 :: "trm" and r2 assume eval1:"[\\<^sub>1]\' \ r1" assume eval2:"[\\<^sub>2]\' \ r2" assume rep:"\ REP \'" assume "(\ REP \' \ ([\\<^sub>1]<>\ \\<^sub>P r1))" then have IH1:"[\\<^sub>1]<>\ \\<^sub>P r1" using rep by auto assume "(\ REP \' \ ([\\<^sub>2]<>\ \\<^sub>P r2))" then have IH2:"[\\<^sub>2]<>\ \\<^sub>P r2" using rep by auto obtain l1 u1 l2 u2 where lu1:"([\\<^sub>1]<> \) = (l1, u1) " and lu2:"([\\<^sub>2]<> \) = (l2, u2)" using IH1 IH2 repP_def by auto have l1:"l1 \\<^sub>L r1" using IH1 lu1 unfolding repP_def by auto have u1:"u1 \\<^sub>U r1" using IH1 lu1 unfolding repP_def by auto have l2:"l2 \\<^sub>L r2" using IH2 lu2 unfolding repP_def by auto have u2:"u2 \\<^sub>U r2" using IH2 lu2 unfolding repP_def by auto then have "([(Times \\<^sub>1 \\<^sub>2)]<>\) = (tl l1 u1 l2 u2, tu l1 u1 l2 u2)" using lu1 lu2 unfolding wTimesU Let_def by auto have lBound:"(tl l1 u1 l2 u2 \\<^sub>L r1 * r2)" using l1 u1 l2 u2 tl_lemma by auto have uBound:"(tu l1 u1 l2 u2 \\<^sub>U r1 * r2)" using l1 u1 l2 u2 tu_lemma by auto have "(tl l1 u1 l2 u2, tu l1 u1 l2 u2) \\<^sub>P (r1 * r2)" unfolding repP_def Let_def using lBound uBound by auto then show "[Times \\<^sub>1 \\<^sub>2]<>\ \\<^sub>P r1 * r2" using lu1 lu2 by auto next case rtsem_Max fix \\<^sub>1 :: "trm" and \' r1 and \\<^sub>2 :: "trm" and r2 assume eval1:"([\\<^sub>1]\' \ r1)" assume eval2:"([\\<^sub>2]\' \ r2)" assume rep:"\ REP \'" assume "(\ REP \' \ [\\<^sub>1]<>\ \\<^sub>P r1)" then have IH1:"[\\<^sub>1]<>\ \\<^sub>P r1" using rep by auto assume "(\ REP \' \ [\\<^sub>2]<>\ \\<^sub>P r2)" then have IH2:"[\\<^sub>2]<>\ \\<^sub>P r2" using rep by auto obtain l1 u1 l2 u2 where lu1:"([\\<^sub>1]<> \) = (l1, u1)" and lu2:"([\\<^sub>2]<> \) = (l2, u2)" using IH1 IH2 repP_def by auto from IH1 IH2 obtain ub1 ub2 lb1 lb2:: real where urep1:"(ub1 \ r1) \ (snd ([\\<^sub>1]<>\) \\<^sub>E ub1)" and urep2:"(ub2 \ r2) \ (snd ([\\<^sub>2]<>\) \\<^sub>E ub2)" and lrep1:"(lb1 \ r1) \ (fst ([\\<^sub>1]<>\) \\<^sub>E lb1)" and lrep2:"(lb2 \ r2) \ (fst ([\\<^sub>2]<>\) \\<^sub>E lb2)" using repP_def repU_def repL_def by auto have lbound:"wmax l1 l2 \\<^sub>L max r1 r2" by (metis dual_order.trans fst_conv le_cases lrep1 lrep2 lu1 lu2 max_def repL_def wmax.elims) have ubound:"wmax u1 u2 \\<^sub>U max r1 r2" by (metis real_max_mono lu1 lu2 repU_def snd_conv urep1 urep2 wmax_lemma) have "([trm.Max \\<^sub>1 \\<^sub>2]<>\) = (wmax l1 l2, wmax u1 u2)" using lu1 lu2 unfolding wMaxU Let_def by (simp) then show "[trm.Max \\<^sub>1 \\<^sub>2]<>\ \\<^sub>P max r1 r2" unfolding repP_def using lbound ubound lu1 lu2 by auto next case rtsem_Min fix \\<^sub>1 :: "trm" and \' r1 and \\<^sub>2 :: "trm" and r2 assume eval1:"([\\<^sub>1]\' \ r1)" assume eval2:"([\\<^sub>2]\' \ r2)" assume rep:"\ REP \'" assume "(\ REP \' \ [\\<^sub>1]<>\ \\<^sub>P r1)" then have IH1:"[\\<^sub>1]<>\ \\<^sub>P r1" using rep by auto assume "(\ REP \' \ [\\<^sub>2]<>\ \\<^sub>P r2)" then have IH2:"[\\<^sub>2]<>\ \\<^sub>P r2" using rep by auto obtain l1 u1 l2 u2 where lu1:"([\\<^sub>1]<> \) = (l1, u1)" and lu2:"([\\<^sub>2]<> \) = (l2, u2)" using IH1 IH2 repP_def by auto from IH1 IH2 obtain ub1 ub2 lb1 lb2:: real where urep1:"(ub1 \ r1) \ (snd ([\\<^sub>1]<>\) \\<^sub>E ub1)" and urep2:"(ub2 \ r2) \ (snd ([\\<^sub>2]<>\) \\<^sub>E ub2)" and lrep1:"(lb1 \ r1) \ (fst ([\\<^sub>1]<>\) \\<^sub>E lb1)" and lrep2:"(lb2 \ r2) \ (fst ([\\<^sub>2]<>\) \\<^sub>E lb2)" using prod.case_eq_if repP_def repU_def repL_def by auto have lbound:"wmin l1 l2 \\<^sub>L min r1 r2" by (metis fst_conv lrep1 lrep2 lu1 lu2 min.mono repL_def wmin_lemma) have ubound:"wmin u1 u2 \\<^sub>U min r1 r2" using lu1 lu2 min_le_iff_disj repU_def urep1 urep2 by auto have "([Min \\<^sub>1 \\<^sub>2]<>\) = (wmin l1 l2, wmin u1 u2)" using lu1 lu2 unfolding wMinU Let_def by auto then show "[Min \\<^sub>1 \\<^sub>2]<>\ \\<^sub>P min r1 r2" unfolding repP_def using lbound ubound lu1 lu2 by auto next case rtsem_Neg fix \ :: "trm" and \' r assume eval:"[\]\' \ r" assume rep:"\ REP \'" assume "(\ REP \' \ [\]<>\ \\<^sub>P r)" then have IH:"[\]<>\ \\<^sub>P r" using rep by auto obtain l1 u1 where lu:"([\]<> \) = (l1, u1)" using IH repP_def by auto from IH obtain ub lb:: real where urep:"(ub \ r) \ (snd ([\]<>\) \\<^sub>E ub)" and lrep:"(lb \ r) \ (fst ([\]<>\) \\<^sub>E lb)" using repP_def repU_def repL_def by auto have ubound:"((wneg u1) \\<^sub>L (uminus r))" by (metis real_minus_le_minus lu repL_def snd_conv urep wneg_lemma) have lbound:"((wneg l1) \\<^sub>U (uminus r))" using real_minus_le_minus lu repU_def lrep wneg_lemma by (metis fst_conv) show "[Neg \]<>\ \\<^sub>P - r" unfolding repP_def Let_def using ubound lbound lu by (auto) next case rtsem_Abs fix \ :: "trm" and \' r assume eval:"[\]\' \ r" assume rep:"\ REP \'" assume "(\ REP \' \ [\]<>\ \\<^sub>P r)" then have IH:"[\]<>\ \\<^sub>P r" using rep by auto obtain l1 u1 where lu:"([\]<> \) = (l1, u1)" using IH repP_def by auto from IH obtain ub lb:: real where urep:"(ub \ r) \ (snd ([\]<>\) \\<^sub>E ub)" and lrep:"(lb \ r) \ (fst ([\]<>\) \\<^sub>E lb)" using prod.case_eq_if repP_def repU_def repL_def by auto have lbound:"wmax l1 (wneg u1) \\<^sub>L (abs r)" apply(simp only: repL_def) apply(rule exI[where x="max lb (- ub)"]) apply(rule conjI) using lrep wmax_lemma lu urep wneg_lemma by auto have ubound:"(wmax u1 (wneg l1) \\<^sub>U (abs r))" apply(simp only: repU_def) apply(rule exI[where x="max ub (- lb)"]) using lrep wmax_lemma lu urep wneg_lemma by auto show "[Abs \]<>\ \\<^sub>P abs r" using repP_def Let_def ubound lbound lu lu wAbsU by auto qed text\Every word represents some real\ lemma word_rep:"\bw::bword. \r::real. Rep_bword bw \\<^sub>E r" proof - fix bw obtain w where weq:"w = Rep_bword bw" by auto have negInfCase:"w = NEG_INF \ ?thesis bw" apply(rule exI[where x="-((2 ^ 31) -1)"]) using weq by (auto simp add: repe.simps) have posInfCase:"w = POS_INF \ ?thesis bw" apply(rule exI[where x="((2 ^ 31) -1)"]) using weq by (auto simp add: repe.simps) have boundU:"w \ NEG_INF \ w \ POS_INF \ sint (Rep_bword bw) < sint POS_INF" using Rep_bword [of bw] by (auto simp: less_le weq [symmetric] dest: sint_dist) have boundL:"w \ NEG_INF \ w \ POS_INF \ sint NEG_INF < sint (Rep_bword bw)" using Rep_bword [of bw] by (auto simp: less_le weq [symmetric] dest: sint_dist) have intCase:"w \ NEG_INF \ w \ POS_INF \ ?thesis bw" apply(rule exI[where x=" (real_of_int (sint (Rep_bword bw)))"]) apply(rule repINT) using boundU boundL by(auto) then show "?thesis bw" apply(cases "w = POS_INF") apply(cases "w = NEG_INF") using posInfCase intCase negInfCase by auto qed text\Every term has a value\ lemma eval_tot:"(\r. ([\::trm]\' \ r))" proof (induction "\") qed (auto simp add: Min_def word_rep bword_neg_one_def, blast?) text\Interval formula semantics soundly implies real semantics\ lemma fml_sound: fixes \::"formula" and \::"wstate" shows "(wfsem \ \ b) \ (\ REP \') \ (rfsem \ \' b)" proof (induction arbitrary: \' rule: wfsem.induct) case (wGreaterT t1 v t2 w) assume wle:"wgreater (fst ([t1]<>v)) (snd ([t2]<>v))" assume rep:"v REP w" obtain r1 and r2 where eval1:"[t1]w \ r1" and eval2:"[t2]w \ r2" using eval_tot[of t1 w] eval_tot[of t2 w] by (auto) note rep1 = trm_sound[of t1 w r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 w r2, where \=v, OF eval2 rep] show "[Greater t1 t2]w \ True" apply(rule rGreaterT[where ?r1.0 = r1, where ?r2.0 = r2]) prefer 3 apply(rule wgreater_lemma[where ?w1.0="fst([t1]<> v)", where ?w2.0="snd([t2]<> v)"]) using rep1 rep2 wle repP_def repL_def repU_def eval1 eval2 by ((simp add: prod.case_eq_if | blast)+) next case (wGreaterF t2 v t1 v') assume wl:"wgeq (fst ([t2]<>v)) (snd ([t1]<>v))" assume rep:"v REP v'" obtain r1 r2:: real where eval1:"(rtsem t1 v' r1)" and eval2:"rtsem t2 v' r2" using eval_tot[of t1 v'] eval_tot[of t2 v'] by (auto) note rep1 = trm_sound[of t1 v' r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 v' r2, where \=v, OF eval2 rep] show "[Greater t1 t2]v' \ False" apply(rule rGreaterF [of t1 v' r1 t2 r2]) apply(rule eval1) apply(rule eval2) apply(rule wgeq_lemma[where ?w1.0="fst([t2]<> v)", where ?w2.0="snd([t1]<> v)"]) using rep1 rep2 repP_def wgeq_lemma wl rep by auto next case (wGeqT t1 v t2 v') assume a1:"wgeq (fst ([t1]<>v)) (snd ([t2]<>v))" assume rep:"v REP v'" obtain r1 r2:: real where eval1:"(rtsem t1 v' r1)" and eval2:"rtsem t2 v' r2" using eval_tot[of t1 v'] eval_tot[of t2 v'] by (auto) note rep1 = trm_sound[of t1 v' r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 v' r2, where \=v, OF eval2 rep] show "[Geq t1 t2]v' \ True" apply(rule rGeqT) apply(rule eval1) apply(rule eval2) using wgeq_lemma eval1 eval2 rep1 rep2 unfolding repP_def Let_def using wgreater_lemma prod.case_eq_if a1 by auto next case (wGeqF t2 v t1 v') assume a1:"wgreater (fst ([t2]<>v)) (snd ([t1]<>v))" assume rep:"v REP v'" obtain r1 r2:: real where eval1:"(rtsem t1 v' r1)" and eval2:"rtsem t2 v' r2" using eval_tot[of t1 v'] eval_tot[of t2 v'] by (auto) note rep1 = trm_sound[of t1 v' r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 v' r2, where \=v, OF eval2 rep] show "[Geq t1 t2]v' \ False" apply(rule rGeqF, rule eval1, rule eval2) using wgeq_lemma eval1 eval2 rep1 rep2 unfolding repP_def Let_def using wgreater_lemma rGreaterF prod.case_eq_if a1 rGreaterF by auto next case (wEqualsT t2 v t1 v') assume eq1:"fst ([t2]<>v) = snd ([t2]<>v)" assume eq2:"snd ([t2]<>v) = snd ([t1]<>v)" assume eq3:"snd ([t1]<>v) = fst ([t1]<>v)" assume rep:"v REP v'" assume neq1:"fst ([t2]<>v) \ NEG_INF" assume neq2:"fst ([t2]<>v) \ POS_INF" obtain r1 r2:: real where eval1:"(rtsem t1 v' r1)" and eval2:"rtsem t2 v' r2" using eval_tot[of t1 v'] eval_tot[of t2 v'] by (auto) note rep1 = trm_sound[of t1 v' r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 v' r2, where \=v, OF eval2 rep] show "[Equals t1 t2]v' \ True" apply(rule rEqualsT, rule eval1, rule eval2) using eq1 eq2 eq3 eval1 eval2 rep1 rep2 unfolding repP_def Let_def repL_def repU_def repe.simps using neq1 neq2 by auto next case (wEqualsF1 t1 v t2 v') assume wle:"wgreater (fst ([t1]<>v)) (snd ([t2]<>v))" assume rep:"v REP v'" obtain r1 r2:: real where eval1:"(rtsem t1 v' r1)" and eval2:"rtsem t2 v' r2" using eval_tot[of t1 v'] eval_tot[of t2 v'] by (auto) note rep1 = trm_sound[of t1 v' r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 v' r2, where \=v, OF eval2 rep] show "[Equals t1 t2]v' \ False" apply(rule rEqualsF, rule eval1, rule eval2) using wgeq_lemma eval1 eval2 rep1 rep2 wgreater_lemma rGreaterF prod.case_eq_if wle unfolding repP_def by (metis (no_types, lifting) less_irrefl) next case (wEqualsF2 t2 v t1 v') assume wle:"wgreater (fst ([t2]<>v)) (snd ([t1]<>v))" assume rep:"v REP v'" obtain r1 r2:: real where eval1:"(rtsem t1 v' r1)" and eval2:"rtsem t2 v' r2" using eval_tot[of t1 v'] eval_tot[of t2 v'] by (auto) note rep1 = trm_sound[of t1 v' r1, where \=v, OF eval1 rep] note rep2 = trm_sound[of t2 v' r2, where \=v, OF eval2 rep] show "[Equals t1 t2]v' \ False" apply(rule rEqualsF, rule eval1, rule eval2) using wgeq_lemma eval1 eval2 rep1 rep2 wgreater_lemma rGreaterF prod.case_eq_if wle unfolding repP_def by (metis (no_types, lifting) less_irrefl) qed (auto) lemma rep_upd:"\ = (\(Inr x := snd([\]<>\)))(Inl x := fst([\]<>\)) \ \ REP \' \ ([\::trm]\' \ r) \ \ REP \'(x := r)" apply(rule REPI) apply(rule conjI) apply(unfold repL_def) using trm_sound prod.case_eq_if repP_def repstate_simps repL_def apply(metis (no_types, lifting) Inl_Inr_False fun_upd_apply sum.inject(1)) using repP_def repstate_simps repU_def apply(auto simp add: repU_def) by (metis (full_types) surjective_pairing trm_sound) text\Interval program semantics soundly contains real semantics existentially\ theorem interval_program_sound: fixes \::"prog" shows "([[\]] \ \ \) \ \ REP \' \ (\\'. (\ REP \') \ ([\] \' \ \'))" proof (induction arbitrary: \' rule: wpsem.induct) case (wTest \ \ \ \') assume sem:"[[\]]\ \ True" and eq:"\ = \" and rep:"\ REP \'" show ?case apply(rule exI[where x=\']) using sem rep by (auto simp add: eq fml_sound rep) next case (wAssign \ \ x \ \') assume eq:"\ = \(Inr x := snd ([\]<>\), Inl x := fst ([\]<>\))" and rep:"\ REP \'" obtain r::real where eval:"([\::trm]\' \ r)" using eval_tot by auto show ?case apply(rule exI[where x="\'(x := r)"]) apply(rule conjI) apply(rule rep_upd[OF eq rep eval]) apply auto apply(rule exI[where x=r]) by (auto simp add: eval) next case (wSeq \ \ \ \ \ \') then show ?case by (simp, blast) next case (wChoice1 a v w b v') then show ?case by auto next case (wChoice2 a v w b v') then show ?case by auto qed end \ No newline at end of file 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,1039 +1,1082 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports "HOL-Library.Word" - Word_Lib More_Arithmetic More_Divides More_Word + Reversed_Bit_Lists 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 [elim?]: 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_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_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 nat_power_less_diff: assumes lt: "(2::nat) ^ n * q < 2 ^ m" shows "q < 2 ^ (m - n)" using lt proof (induct n arbitrary: m) case 0 then show ?case by simp next case (Suc n) have ih: "\m. 2 ^ n * q < 2 ^ m \ q < 2 ^ (m - n)" and prem: "2 ^ Suc n * q < 2 ^ m" by fact+ show ?case proof (cases m) case 0 then show ?thesis using Suc by simp next case (Suc m') then show ?thesis using prem by (simp add: ac_simps ih) qed qed 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 is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" apply (simp add: is_aligned_mask eq_zero_set_bl) apply (clarsimp simp: in_set_conv_nth word_size) apply (simp add: to_bl_nth word_size cong: conj_cong) apply (simp add: diff_diff_less) apply safe apply (case_tac "n \ LENGTH('a)") prefer 2 apply (rule_tac x=i in exI) apply clarsimp apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") apply (erule exE) apply (rule_tac x=j in exI) apply clarsimp apply (thin_tac "w !! t" for t) apply (rule_tac x="i + n - LENGTH('a)" in exI) apply clarsimp apply arith apply (rule_tac x="LENGTH('a) - n + i" in exI) apply clarsimp apply arith done -lemma power_overflow: - "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" - by simp - lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" proof - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ to_bl (2 ^ n * (of_nat q :: 'a word)) = drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) show ?thesis using aligned by (auto simp: rl elim: is_alignedE) qed lemma is_aligned_drop: fixes w::"'a::len word" assumes "is_aligned w n" "n \ LENGTH('a)" shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" proof - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" by (rule is_aligned_replicate) fact+ then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp also have "\ = replicate n False" by simp finally show ?thesis . qed lemma less_is_drop_replicate: fixes x::"'a::len word" assumes lt: "x < 2 ^ n" shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" by (metis assms bl_and_mask' less_mask_eq) lemma is_aligned_add_conv: fixes off::"'a::len word" assumes aligned: "is_aligned w n" and offv: "off < 2 ^ n" shows "to_bl (w + off) = (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" proof cases assume nv: "n \ LENGTH('a)" show ?thesis proof (subst aligned_bl_add_size, simp_all only: word_size) show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) from offv show "take (LENGTH('a) - n) (to_bl off) = replicate (LENGTH('a) - n) False" by (subst less_is_drop_replicate, assumption) simp qed fact next assume "\ n \ LENGTH('a)" with offv show ?thesis by (simp add: power_overflow) qed 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) 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) 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 (k << m) 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 (k << m)" 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 (k << m) = unat (2 ^ m * k)" by (simp add: shiftl_t2n) 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 (k << m) = 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 shiftl_zero_size 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 is_aligned_replicateI: "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" apply (simp add: is_aligned_to_bl word_size) apply (subgoal_tac "length addr = LENGTH('a) - n") apply (simp add: replicate_not_True) apply (drule arg_cong [where f=length]) apply simp done lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" apply (subst shiftl_1 [symmetric]) apply (subst bl_shiftl) apply (simp add: to_bl_1 min_def word_size) done 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 xor_2p_to_bl: fixes x::"'a::len word" shows "to_bl (x XOR 2^n) = (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" proof - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" by simp show ?thesis apply simp apply (rule conjI) apply (clarsimp simp: word_size) apply (simp add: bl_word_xor to_bl_2p) apply (subst x) apply (subst zip_append) apply simp apply (simp add: map_zip_replicate_False_xor drop_minus) apply (auto simp add: word_size nth_w2p intro!: word_eqI) done qed lemma aligned_add_xor: assumes al: "is_aligned (x::'a::len word) n'" and le: "n < n'" shows "(x + 2^n) XOR 2^n = x" proof cases assume "n' < LENGTH('a)" with assms show ?thesis apply - apply (rule word_bl.Rep_eqD) apply (subst xor_2p_to_bl) apply simp apply (subst is_aligned_add_conv, simp, simp add: word_less_nat_alt)+ apply (simp add: to_bl_2p nth_append) apply (cases "n' = Suc n") apply simp apply (subst is_aligned_replicate [where n="Suc n", simplified, symmetric]; simp) apply (subgoal_tac "\ LENGTH('a) - Suc n \ LENGTH('a) - n'") prefer 2 apply arith apply (subst replicate_Suc [symmetric]) apply (subst replicate_add [symmetric]) apply (simp add: is_aligned_replicate [simplified, symmetric]) done next assume "\ n' < LENGTH('a)" show ?thesis using al apply (rule is_alignedE) using \\ n' < LENGTH('a)\ by auto qed lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ \ \xs. to_bl w = xs @ replicate n False \ length xs = size w - n" apply (subst is_aligned_replicate, assumption+) apply (rule exI, rule conjI, rule refl) apply (simp add: word_size) done 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" apply (erule is_aligned_get_word_bits) apply (frule is_aligned_add_conv [rotated, where w=0]) apply (simp add: is_aligned_iff_dvd_nat) apply simp apply (drule is_aligned_replicateD) apply simp apply (clarsimp simp: word_size) apply (subst (asm) replicate_add [symmetric]) apply (drule arg_cong[where f="of_bl :: bool list \ 'a::len word"]) apply simp apply (simp only: replicate.simps[symmetric, where x=False] drop_replicate) done 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_nth: "is_aligned p m = (\n < m. \p !! n)" apply (clarsimp simp: is_aligned_mask bang_eq word_size) apply (rule iffI) apply clarsimp apply (case_tac "n < size p") apply (simp add: word_size) apply (drule test_bit_size) apply simp apply clarsimp 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 (metis and_not_mask is_aligned_shift is_aligned_weaken) 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) \ \ 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 apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_less test_bit_eq_bit) apply (metis is_aligned_nth not_le test_bit_eq_bit) done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (w << m) n" by (simp add: is_aligned_nth nth_shiftl) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (w >> m) n" by (simp add: is_aligned_nth nth_shiftr) lemma is_aligned_shiftl_self: "is_aligned (p << n) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p AND NOT (mask n) = p" by (metis add.left_neutral is_aligned_mask word_plus_and_or_coroll2) lemma is_aligned_shiftr_shiftl: "is_aligned w n \ w >> n << n = w" by (metis and_not_mask is_aligned_neg_mask_eq) lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ ((x >> n) AND mask v) << n = x AND mask (v + n)" apply (rule word_eqI) apply (simp add: word_size nth_shiftl nth_shiftr) apply (subgoal_tac "\m. 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) 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 (simp add: and_not_mask is_aligned_shift) 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]) 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 (simp add: word_size_bl) 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: + "NOT(mask n) \ w \ (\i \ {n ..< size w}. w !! i)" + for w :: \'a::len word\ + by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_test_bit) + 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,265 +1,263 @@ (* * 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 begin -declare of_nat_diff [simp] - 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) + 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) = max_word" 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 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,822 +1,871 @@ section \Lemmas on words\ theory More_Word imports "HOL-Library.Word" More_Arithmetic More_Divides begin lemma Suc_0_lt_2p_len_of: "Suc 0 < 2 ^ LENGTH('a :: len)" by (metis One_nat_def len_gt_0 lessI numeral_2_eq_2 one_less_power) 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 shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" apply (induct x) apply simp apply (simp add: shiftl1_2t) 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 le_shiftr1: \shiftr1 u \ shiftr1 v\ if \u \ v\ using that proof transfer fix k l :: int assume \take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ then have \take_bit LENGTH('a) (drop_bit 1 (take_bit LENGTH('a) k)) \ take_bit LENGTH('a) (drop_bit 1 (take_bit LENGTH('a) l))\ apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done then show \take_bit LENGTH('a) (take_bit LENGTH('a) k div 2) \ take_bit LENGTH('a) (take_bit LENGTH('a) l div 2)\ by (simp add: drop_bit_eq_div) qed 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 = max_word \ 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 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 \ max_word \ (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 \ max_word \ 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 \ max_word \ 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 \ max_word" 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) + end diff --git a/thys/Word_Lib/More_Word_Operations.thy b/thys/Word_Lib/More_Word_Operations.thy --- a/thys/Word_Lib/More_Word_Operations.thy +++ b/thys/Word_Lib/More_Word_Operations.thy @@ -1,646 +1,651 @@ section \Misc word operations\ theory More_Word_Operations imports "HOL-Library.Word" Reversed_Bit_Lists Aligned More_Divides More_Misc "HOL-Library.Sublist" begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 AND NOT (2 ^ n - 1)" (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if w !! n then w OR NOT (mask n) else w AND mask n" lemma sign_extend_eq_signed_take_bit: \sign_extend = signed_take_bit\ proof (rule ext)+ fix n and w :: \'a::len word\ show \sign_extend n w = signed_take_bit n w\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ then show \bit (sign_extend n w) q \ bit (signed_take_bit n w) q\ by (auto simp add: test_bit_eq_bit bit_signed_take_bit_iff sign_extend_def bit_and_iff bit_or_iff bit_not_iff bit_mask_iff not_less exp_eq_0_imp_not_bit not_le min_def) qed qed definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def is_aligned_mask mask_eq_decr_exp word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis mask_eq_decr_exp is_aligned_add_helper p_assoc_help power_2_ge_iff) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst mask_eq_decr_exp [symmetric]) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (simp flip: drop_bit_eq_div unat_drop_bit_eq) (metis leI le_unat_uoi unat_mono) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 AND NOT (mask sz)" by (simp add: alignUp_def flip: mask_eq_decr_exp) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 AND NOT (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q AND NOT (mask sz)) = (p - ((alignUp q sz) AND NOT (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) apply (simp add: alignUp_distance is_aligned_neg_mask_eq mask_out_add_aligned) apply (metis add_cancel_right_right alignUp_distance and_mask_eq_iff_le_mask word_plus_and_or_coroll2) done lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) +lemma sign_extend_bitwise_if': + \i < LENGTH('a) \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)\ + for w :: \'a::len word\ + using sign_extend_bitwise_if [of i w e] by (simp add: word_size) + lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w OR NOT (mask (Suc n)) else w AND mask (Suc n))" by (rule bit_word_eqI) (auto simp add: bit_simps sign_extend_eq_signed_take_bit min_def test_bit_eq_bit less_Suc_eq_le) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if) lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply auto apply (auto simp add: bit_eq_iff) apply (simp_all add: bit_simps sign_extend_eq_signed_take_bit not_le min_def sign_extended_def test_bit_eq_bit word_size split: if_splits) using le_imp_less_or_eq apply auto[1] apply (metis bit_imp_le_length nat_less_le) apply (metis Suc_leI Suc_n_not_le_n le_trans nat_less_le) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by (rule bit_word_eqI) (simp add: sign_extend_eq_signed_take_bit bit_simps) lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w AND mask (Suc n) = v AND mask (Suc n) \ sign_extend n w = sign_extend n v" by (simp flip: take_bit_eq_mask add: sign_extend_eq_signed_take_bit signed_take_bit_eq_iff_take_bit_eq) lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f AND mask e = f" by (fastforce intro: subst[where P="\f. f AND mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr AND NOT (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) definition "limited_and (x :: 'a :: len word) y \ (x AND y = x)" lemma limited_and_eq_0: "\ limited_and x z; y AND NOT z = y \ \ x AND y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(AND)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y AND z = z \ \ x AND y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_eq_decr_exp, folded limited_and_def] lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] not_one shiftl_shiftr1[unfolded word_size mask_eq_decr_exp] shiftl_shiftr2[unfolded word_size mask_eq_decr_exp] definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_eq: \from_bool = of_bool\ by (simp add: fun_eq_iff from_bool_def) lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x AND 1) = (x !! 0)" by (simp add: test_bit_word_eq to_bool_def and_one_eq mod_2_eq_odd) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) AND 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) end \ No newline at end of file diff --git a/thys/Word_Lib/Rsplit.thy b/thys/Word_Lib/Rsplit.thy --- a/thys/Word_Lib/Rsplit.thy +++ b/thys/Word_Lib/Rsplit.thy @@ -1,148 +1,160 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA *) theory Rsplit imports "HOL-Library.Word" Bits_Int begin definition word_rsplit :: "'a::len word \ 'b::len word list" where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = map word_of_int (bin_rsplit (LENGTH('a::len)) (LENGTH('b), take_bit (LENGTH('b)) (numeral bin)))" by (simp add: word_rsplit_def of_nat_take_bit) lemmas word_rsplit_no_cl [simp] = word_rsplit_no [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] text \ This odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word.\ lemma word_rsplit_same: "word_rsplit w = [w]" apply (simp add: word_rsplit_def bin_rsplit_all) apply transfer apply simp done lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def split: prod.split) lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) apply (rule_tac f = "\x. bin_nth x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) apply simp_all apply (simp add : word_size rev_map) apply (rule trans) defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) apply simp using bin_rsplit_size_sign take_bit_int_eq_self_iff by blast lemma test_bit_rsplit_alt: \(word_rsplit w :: 'b::len word list) ! i !! m \ w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ for w :: \'a::len word\ apply (rule trans) apply (rule test_bit_cong) apply (rule rev_nth [of _ \rev (word_rsplit w)\, simplified rev_rev_ident]) apply simp apply (rule that(1)) apply simp apply (rule test_bit_rsplit) apply (rule refl) apply (rule asm_rl) apply (rule that(2)) apply (rule diff_Suc_less) apply (rule that(3)) done lemma word_rsplit_len_indep [OF refl refl refl refl]: "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ word_rsplit v = sv \ length su = length sv" by (auto simp: word_rsplit_def bin_rsplit_len_indep) lemma length_word_rsplit_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) lemmas length_word_rsplit_lt_size = length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] lemma length_word_rsplit_exp_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" by (auto simp: word_rsplit_def word_size bin_rsplit_len) lemma length_word_rsplit_even_size: "n = LENGTH('a::len) \ size w = m * n \ length (word_rsplit w :: 'a word list) = m" by (cases \LENGTH('a)\) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] \ \alternative proof of \word_rcat_rsplit\\ lemmas tdle = times_div_less_eq_dividend lemmas dtle = xtrans(4) [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) apply (clarsimp simp: test_bit_rcat word_size) apply (subst refl [THEN test_bit_rsplit]) apply (simp_all add: word_size refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) apply safe apply (erule xtrans(7), rule dtle)+ done lemma size_word_rsplit_rcat_size: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ length (word_rsplit frcw::'a word list) = length ws" for ws :: "'a::len word list" and frcw :: "'b::len word" by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) lemma msrevs: "0 < n \ (k * n + m) div n = m div n + k" "(k * n + m) mod n = m mod n" for n :: nat by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" for ws :: "'a::len word list" apply (frule size_word_rsplit_rcat_size, assumption) apply (clarsimp simp add : word_size) apply (rule nth_equalityI, assumption) apply clarsimp apply (rule word_eqI [rule_format]) apply (rule trans) apply (rule test_bit_rsplit_alt) apply (clarsimp simp: word_size)+ apply (rule trans) apply (rule test_bit_rcat [OF refl refl]) apply (simp add: word_size) apply (subst rev_nth) apply arith apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) apply safe apply (simp add: diff_mult_distrib) apply (cases "size ws") apply simp_all done +lemma word_rsplit_upt: + "\ size x = LENGTH('a :: len) * n; n \ 0 \ + \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" + apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") + apply (rule nth_equalityI, simp) + apply (intro allI word_eqI impI) + apply (simp add: test_bit_rsplit_alt word_size) + apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) + apply (simp add: length_word_rsplit_exp_size) + apply (metis mult.commute given_quot_alt word_size word_size_gt_0) + done + end \ No newline at end of file diff --git a/thys/Word_Lib/Signed_Words.thy b/thys/Word_Lib/Signed_Words.thy --- a/thys/Word_Lib/Signed_Words.thy +++ b/thys/Word_Lib/Signed_Words.thy @@ -1,40 +1,66 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Signed Words" theory Signed_Words imports "HOL-Library.Word" begin text \Signed words as separate (isomorphic) word length class. Useful for tagging words in C.\ typedef ('a::len0) signed = "UNIV :: 'a set" .. lemma card_signed [simp]: "CARD (('a::len0) signed) = CARD('a)" unfolding type_definition.card [OF type_definition_signed] by simp instantiation signed :: (len0) len0 begin definition len_signed [simp]: "len_of (x::'a::len0 signed itself) = LENGTH('a)" instance .. end instance signed :: (len) len by (intro_classes, simp) +lemma scast_scast_id [simp]: + "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" + "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" + by (auto simp: is_up scast_up_scast_id) + +lemma ucast_scast_id [simp]: + "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" + by transfer (simp add: take_bit_signed_take_bit) + +lemma scast_of_nat [simp]: + "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" + by transfer (simp add: take_bit_signed_take_bit) + +lemma scast_ucast_id [simp]: + "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" + by transfer (simp add: take_bit_signed_take_bit) + +lemma scast_eq_scast_id [simp]: + "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" + by (metis ucast_scast_id) + +lemma ucast_eq_ucast_id [simp]: + "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" + by (metis scast_ucast_id) + +lemma scast_ucast_norm [simp]: + "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" + "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" + by (metis scast_ucast_id ucast_scast_id)+ + type_synonym 'a sword = "'a signed word" -type_synonym sword8 = "8 sword" -type_synonym sword16 = "16 sword" -type_synonym sword32 = "32 sword" -type_synonym sword64 = "64 sword" end diff --git a/thys/Word_Lib/Traditional_Infix_Syntax.thy b/thys/Word_Lib/Traditional_Infix_Syntax.thy --- a/thys/Word_Lib/Traditional_Infix_Syntax.thy +++ b/thys/Word_Lib/Traditional_Infix_Syntax.thy @@ -1,855 +1,872 @@ (* Author: Jeremy Dawson, NICTA *) section \Operation variants with traditional syntax\ theory Traditional_Infix_Syntax imports "HOL-Library.Word" More_Word begin class semiring_bit_syntax = semiring_bit_shifts begin definition test_bit :: \'a \ nat \ bool\ (infixl "!!" 100) where test_bit_eq_bit: \test_bit = bit\ definition shiftl :: \'a \ nat \ 'a\ (infixl "<<" 55) where shiftl_eq_push_bit: \a << n = push_bit n a\ definition shiftr :: \'a \ nat \ 'a\ (infixl ">>" 55) where shiftr_eq_drop_bit: \a >> n = drop_bit n a\ end instance word :: (len) semiring_bit_syntax .. context includes lifting_syntax begin lemma test_bit_word_transfer [transfer_rule]: \(pcr_word ===> (=)) (\k n. n < LENGTH('a) \ bit k n) (test_bit :: 'a::len word \ _)\ by (unfold test_bit_eq_bit) transfer_prover lemma shiftl_word_transfer [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (\k n. push_bit n k) shiftl\ by (unfold shiftl_eq_push_bit) transfer_prover lemma shiftr_word_transfer [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (\k n. (drop_bit n \ take_bit LENGTH('a)) k) (shiftr :: 'a::len word \ _)\ by (unfold shiftr_eq_drop_bit) transfer_prover end lemma test_bit_word_eq: \test_bit = (bit :: 'a::len word \ _)\ by (fact test_bit_eq_bit) lemma shiftl_word_eq: \w << n = push_bit n w\ for w :: \'a::len word\ by (fact shiftl_eq_push_bit) lemma shiftr_word_eq: \w >> n = drop_bit n w\ for w :: \'a::len word\ by (fact shiftr_eq_drop_bit) lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" for u v :: "'a::len word" by (simp add: bit_eq_iff test_bit_eq_bit fun_eq_iff) lemma test_bit_size: "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 \ u !! n = v !! n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ u !! x = v !! x" for u v :: "'a::len word" by simp lemma test_bit_bin': "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: \test_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) (test_bit :: 'a::len word \ _)" by (simp only: test_bit_eq_bit) transfer_prover lemma test_bit_wi [simp]: "(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 \ (x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n) \ (x XOR y) !! n = (x !! n \ y !! n) \ (NOT x) !! n = (\ 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: "(x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n)" for x :: "'a::len word" by transfer (auto simp add: bit_or_iff bit_and_iff) 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_numeral [simp]: "(numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bit (numeral w :: int) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bit (- numeral w :: int) n" by transfer (rule refl) lemma test_bit_1 [iff]: "(1 :: 'a::len word) !! n \ n = 0" by transfer (auto simp add: bit_1_iff) lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" by transfer simp lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \ n < LENGTH('a)" by transfer simp lemma shiftl1_code [code]: \shiftl1 w = push_bit 1 w\ by transfer (simp add: ac_simps) 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 shiftr1_code [code]: \shiftr1 w = drop_bit 1 w\ by transfer (simp add: drop_bit_Suc) lemma shiftl_def: \w << n = (shiftl1 ^^ n) w\ 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 = (shiftr1 ^^ n) w\ proof - have \shiftr1 ^^ n = (drop_bit n :: 'a word \ 'a word)\ apply (induction n) apply simp apply (simp only: shiftr1_eq_div_2 [abs_def] drop_bit_eq_div [abs_def] funpow_Suc_right) apply (use div_exp_eq [of _ 1, where ?'a = \'a word\] in simp) done then show ?thesis by (simp add: shiftr_eq_drop_bit) qed 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: shiftl_word_eq bit_push_bit_iff exp_eq_zero_iff not_le) lemma bit_shiftr_word_iff [bit_simps]: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: shiftr_word_eq bit_drop_bit_eq) lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k))\ by (simp flip: signed_take_bit_decr_length_iff) lemma sshiftr_eq [code]: \w >>> n = signed_drop_bit n w\ by transfer simp lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (sshiftr1 ^^ n) w\ apply (rule sym) apply (simp add: sshiftr1_eq_signed_drop_bit_Suc_0 sshiftr_eq) 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 sshift1_code [code]: \sshiftr1 w = signed_drop_bit 1 w\ by transfer (simp add: drop_bit_Suc) lemma sshiftr_0 [simp]: "0 >>> n = 0" by transfer simp lemma sshiftr_n1 [simp]: "-1 >>> n = -1" by transfer simp lemma bit_sshiftr_word_iff [bit_simps]: \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\ apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma nth_sshiftr : "(w >>> m) !! n = (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le ac_simps) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma sshiftr_numeral [simp]: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ apply (rule word_eqI) apply (cases \LENGTH('a)\) apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr bit_signed_take_bit_iff min_def not_le not_less less_Suc_eq_le ac_simps) done setup \ Context.theory_map (fold SMT_Word.add_word_shift' [ (\<^term>\shiftl :: 'a::len word \ _\, "bvshl"), (\<^term>\shiftr :: 'a::len word \ _\, "bvlshr"), (\<^term>\sshiftr :: 'a::len word \ _\, "bvashr") ]) \ 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 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 sshiftr_eq) 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 w !! n else 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: "(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: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" by transfer (auto simp add: bit_exp_iff) lemma bang_is_le: "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 lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" by transfer simp lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" by transfer simp lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" by transfer (auto simp add: bit_double_iff) lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ 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_shiftr1: "shiftr1 w !! n = w !! Suc n" by transfer (auto simp add: bit_take_bit_iff simp flip: bit_Suc) lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct "m" arbitrary: n) apply (auto simp add: nth_shiftr1) done lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" apply transfer apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done 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 (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) 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) lemma shiftl_numeral [simp]: \numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\ by (fact shiftl_word_eq) 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 (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr_numeral [simp]: \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ by (fact shiftr_word_eq) lemma nth_mask [simp]: \(mask n :: 'a::len word) !! i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) 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 nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \ m < LENGTH('a))" by (simp add: slice_shiftr nth_ucast nth_shiftr) 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!: *) done then show ?thesis apply (simp only: One_nat_def word_size shiftr_word_eq 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 test_bit_cat [OF refl]: "wc = word_cat a b \ wc !! n = (n < size wc \ (if n < size b then b !! n else 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. b !! n = (n < size b \ c !! n) \ a !! m = (m < size a \ c !! (m + size b)))" by (auto simp add: word_split_bin' test_bit_bin bit_unsigned_iff word_size bit_drop_bit_eq ac_simps exp_eq_zero_iff dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. b !! n \ n < size b \ c !! n) \ (\m::nat. a !! m \ m < size a \ c !! (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. b !! n = (n < size b \ c !! n)) \ (\m::nat. a !! m = (m < size a \ 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 \ rc !! n = (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) (simp add: test_bit_eq_bit) lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" by (fact nth_minus1) lemma shiftr_x_0 [iff]: "x >> 0 = x" for x :: "'a::len word" by transfer simp lemma shiftl_x_0 [simp]: "x << 0 = x" for x :: "'a::len word" by (simp add: shiftl_t2n) lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_t2n) lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" by (induct xs) auto lemma word_and_1: "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) lemma test_bit_1' [simp]: "(1 :: 'a :: len word) !! n \ 0 < LENGTH('a) \ n = 0" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma word_ops_nth [simp]: fixes x y :: \'a::len word\ shows word_or_nth: "(x OR y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x AND y) !! n = (x !! n \ y !! n)" and word_xor_nth: "(x XOR y) !! n = (x !! n \ y !! n)" by ((cases "n < size x", auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) 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 (unfold shiftr_def) apply (induct_tac "n") apply auto apply (erule le_shiftr1) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma word_leI: "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" apply (rule xtrans(4)) apply (rule word_and_le2) apply (rule word_eqI) apply (simp add: word_ao_nth) apply safe apply assumption apply (erule_tac [2] asm_rl) apply (unfold word_size) by auto 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) 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 (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce lemma shiftl_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: word_ao_nth nth_shiftl, safe) done lemma shiftr_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_shiftr word_ao_nth) done 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)" apply(rule word_eqI) apply(simp add:nth_shiftl word_ao_nth, safe) done lemma shiftr_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_shiftr word_ao_nth) done 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:nth_shiftr nth_shiftl word_size word_ao_nth) 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) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done 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:nth_shiftr nth_shiftl word_size word_ops_nth_size) 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)" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply arith done 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 neg_mask_test_bit: "(NOT(mask n) :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" by (metis not_le nth_mask test_bit_bin word_ops_nth_size word_size) lemma upper_bits_unset_is_l2p: \(\n' \ n. n' < LENGTH('a) \ \ 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 add: test_bit_eq_bit) 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) \ \ 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) \ (x !! n) = False" + by transfer auto + +lemma le_mask_high_bits: + "w \ mask n \ (\i \ {n ..< size w}. \ w !! i)" + for w :: \'a::len word\ + by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) + +lemma test_bit_conj_lt: + "(x !! m \ m < LENGTH('a)) = x !! m" for x :: "'a :: len word" + using test_bit_bin by blast + +lemma neg_test_bit: + "(NOT x) !! n = (\ 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) + end diff --git a/thys/Word_Lib/Word_16.thy b/thys/Word_Lib/Word_16.thy --- a/thys/Word_Lib/Word_16.thy +++ b/thys/Word_Lib/Word_16.thy @@ -1,23 +1,25 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 16" theory Word_16 imports Word_Lemmas Word_Syntax begin type_synonym word16 = "16 word" lemma len16: "len_of (x :: 16 itself) = 16" by simp +type_synonym sword16 = "16 sword" + lemma word16_and_max_simp: \x AND 0xFFFF = x\ for x :: \16 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) end diff --git a/thys/Word_Lib/Word_32.thy b/thys/Word_Lib/Word_32.thy --- a/thys/Word_Lib/Word_32.thy +++ b/thys/Word_Lib/Word_32.thy @@ -1,342 +1,345 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 32" theory Word_32 imports Word_Lemmas Word_8 Word_16 Word_Syntax + Rsplit Bitwise begin type_synonym word32 = "32 word" lemma len32: "len_of (x :: 32 itself) = 32" by simp +type_synonym sword32 = "32 sword" + type_synonym machine_word_len = 32 type_synonym machine_word = "machine_word_len word" definition word_bits :: nat where "word_bits = LENGTH(machine_word_len)" text \The following two are numerals so they can be used as nats and words.\ definition word_size_bits :: "'a :: numeral" where "word_size_bits = 2" definition word_size :: "'a :: numeral" where "word_size = 4" lemma word_bits_conv[code]: "word_bits = 32" unfolding word_bits_def by simp lemma word_size_word_size_bits: "(word_size::nat) = 2 ^ word_size_bits" unfolding word_size_def word_size_bits_def by simp lemma word_bits_word_size_conv: "word_bits = word_size * 8" unfolding word_bits_def word_size_def by simp lemma ucast_8_32_inj: "inj (ucast :: 8 word \ 32 word)" by (rule down_ucast_inj) (clarsimp simp: is_down_def target_size source_size) lemma upto_2_helper: "{0..<2 :: 32 word} = {0, 1}" by (safe; simp) unat_arith lemmas upper_bits_unset_is_l2p_32 = upper_bits_unset_is_l2p [where 'a=32, folded word_bits_def] lemmas le_2p_upper_bits_32 = le_2p_upper_bits [where 'a=32, folded word_bits_def] lemmas le2p_bits_unset_32 = le2p_bits_unset[where 'a=32, folded word_bits_def] lemma word_bits_len_of: "len_of TYPE (32) = word_bits" by (simp add: word_bits_conv) lemmas unat_power_lower32' = unat_power_lower[where 'a=32] lemmas unat_power_lower32 [simp] = unat_power_lower32'[unfolded word_bits_len_of] lemmas word32_less_sub_le' = word_less_sub_le[where 'a = 32] lemmas word32_less_sub_le[simp] = word32_less_sub_le' [folded word_bits_def] lemma word_bits_size: "size (w::word32) = word_bits" by (simp add: word_bits_def word_size) lemmas word32_power_less_1' = word_power_less_1[where 'a = 32] lemmas word32_power_less_1[simp] = word32_power_less_1'[folded word_bits_def] lemma of_nat32_0: "\of_nat n = (0::word32); n < 2 ^ word_bits\ \ n = 0" by (erule of_nat_0, simp add: word_bits_def) lemma unat_mask_2_less_4: "unat (p && mask 2 :: word32) < 4" apply (rule unat_less_helper) apply (rule order_le_less_trans, rule word_and_le1) apply (simp add: mask_eq) done lemmas unat_of_nat32' = unat_of_nat_eq[where 'a=32] lemmas unat_of_nat32 = unat_of_nat32'[unfolded word_bits_len_of] lemmas word_power_nonzero_32 = word_power_nonzero [where 'a=32, folded word_bits_def] lemmas unat_mult_simple = iffD1 [OF unat_mult_lem [where 'a = 32, unfolded word_bits_len_of]] lemmas div_power_helper_32 = div_power_helper [where 'a=32, folded word_bits_def] lemma n_less_word_bits: "(n < word_bits) = (n < 32)" by (simp add: word_bits_def) lemmas of_nat_less_pow_32 = of_nat_power [where 'a=32, folded word_bits_def] lemma lt_word_bits_lt_pow: "sz < word_bits \ sz < 2 ^ word_bits" by (simp add: word_bits_conv) lemma unat_less_word_bits: fixes y :: word32 shows "x < unat y \ x < 2 ^ word_bits" unfolding word_bits_def by (rule order_less_trans [OF _ unat_lt2p]) lemmas unat_mask_word32' = unat_mask[where 'a=32] lemmas unat_mask_word32 = unat_mask_word32'[folded word_bits_def] lemma unat_less_2p_word_bits: "unat (x :: 32 word) < 2 ^ word_bits" apply (simp only: word_bits_def) apply (rule unat_lt2p) done lemma Suc_unat_mask_div: "Suc (unat (mask sz div word_size::word32)) = 2 ^ (min sz word_bits - 2)" apply (case_tac "sz < word_bits") apply (case_tac "2 \ sz") apply (clarsimp simp: word_size_def word_bits_def min_def mask_eq) apply (drule (2) Suc_div_unat_helper [where 'a=32 and sz=sz and us=2, simplified, symmetric]) apply (simp add: not_le word_size_def word_bits_def) apply (case_tac sz, simp add: unat_word_ariths) apply (case_tac nat, simp add: unat_word_ariths unat_mask_word32 min_def word_bits_def) apply simp apply (simp add: unat_word_ariths unat_mask_word32 min_def word_bits_def word_size_def) done lemmas word32_minus_one_le' = word_minus_one_le[where 'a=32] lemmas word32_minus_one_le = word32_minus_one_le'[simplified] lemma ucast_not_helper: fixes a::word8 assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::word32)" proof assume "ucast a = (0xFF::word32)" also have "(0xFF::word32) = ucast (0xFF::word8)" by simp finally show False using a apply - apply (drule up_ucast_inj, simp) apply simp done qed lemma less_4_cases: "(x::word32) < 4 \ x=0 \ x=1 \ x=2 \ x=3" apply clarsimp apply (drule word_less_cases, erule disjE, simp, simp)+ done lemma unat_ucast_8_32: fixes x :: "word8" shows "unat (ucast x :: word32) = unat x" by transfer simp lemma if_then_1_else_0: "((if P then 1 else 0) = (0 :: word32)) = (\ P)" by simp lemma if_then_0_else_1: "((if P then 0 else 1) = (0 :: word32)) = (P)" by simp lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 lemma ucast_le_ucast_8_32: "(ucast x \ (ucast y :: word32)) = (x \ (y :: word8))" by (simp add: ucast_le_ucast) lemma in_16_range: "0 \ S \ r \ (\x. r + x * (16 :: word32)) ` S" "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: word32. r + x * 16) ` S" by (clarsimp simp: image_def elim!: bexI[rotated])+ lemma eq_2_32_0: "(2 ^ 32 :: word32) = 0" by simp lemma x_less_2_0_1: fixes x :: word32 shows "x < 2 \ x = 0 \ x = 1" by (rule x_less_2_0_1') auto lemmas mask_32_max_word = max_word_mask [symmetric, where 'a=32, simplified] lemma of_nat32_n_less_equal_power_2: "n < 32 \ ((of_nat n)::32 word) < 2 ^ n" by (rule of_nat_n_less_equal_power_2, clarsimp simp: word_size) lemma word_rsplit_0: "word_rsplit (0 :: word32) = [0, 0, 0, 0 :: word8]" by (simp add: word_rsplit_def bin_rsplit_def) lemma unat_ucast_10_32 : fixes x :: "10 word" shows "unat (ucast x :: word32) = unat x" by transfer simp lemma bool_mask [simp]: fixes x :: word32 shows "(0 < x && 1) = (x && 1 = 1)" by (rule bool_mask') auto lemma word32_bounds: "- (2 ^ (size (x :: word32) - 1)) = (-2147483648 :: int)" "((2 ^ (size (x :: word32) - 1)) - 1) = (2147483647 :: int)" "- (2 ^ (size (y :: 32 signed word) - 1)) = (-2147483648 :: int)" "((2 ^ (size (y :: 32 signed word) - 1)) - 1) = (2147483647 :: int)" by (simp_all add: word_size) lemma word_ge_min:"sint (x::32 word) \ -2147483648" by (metis sint_ge word32_bounds(1) word_size) lemmas signed_arith_ineq_checks_to_eq_word32' = signed_arith_ineq_checks_to_eq[where 'a=32] signed_arith_ineq_checks_to_eq[where 'a="32 signed"] lemmas signed_arith_ineq_checks_to_eq_word32 = signed_arith_ineq_checks_to_eq_word32' [unfolded word32_bounds] lemmas signed_mult_eq_checks32_to_64' = signed_mult_eq_checks_double_size[where 'a=32 and 'b=64] signed_mult_eq_checks_double_size[where 'a="32 signed" and 'b=64] lemmas signed_mult_eq_checks32_to_64 = signed_mult_eq_checks32_to_64'[simplified] lemmas sdiv_word32_max' = sdiv_word_max [where 'a=32] sdiv_word_max [where 'a="32 signed"] lemmas sdiv_word32_max = sdiv_word32_max'[simplified word_size, simplified] lemmas sdiv_word32_min' = sdiv_word_min [where 'a=32] sdiv_word_min [where 'a="32 signed"] lemmas sdiv_word32_min = sdiv_word32_min' [simplified word_size, simplified] lemmas sint32_of_int_eq' = sint_of_int_eq [where 'a=32] lemmas sint32_of_int_eq = sint32_of_int_eq' [simplified] lemma ucast_of_nats [simp]: "(ucast (of_nat x :: word32) :: sword32) = (of_nat x)" "(ucast (of_nat x :: word32) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word32) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word8) :: sword8) = (of_nat x)" by (simp_all add: of_nat_take_bit take_bit_word_eq_self) lemmas signed_shift_guard_simpler_32' = power_strict_increasing_iff[where b="2 :: nat" and y=31] lemmas signed_shift_guard_simpler_32 = signed_shift_guard_simpler_32'[simplified] lemma word32_31_less: "31 < len_of TYPE (32 signed)" "31 > (0 :: nat)" "31 < len_of TYPE (32)" "31 > (0 :: nat)" by auto lemmas signed_shift_guard_to_word_32 = signed_shift_guard_to_word[OF word32_31_less(1-2)] signed_shift_guard_to_word[OF word32_31_less(3-4)] lemma le_step_down_word_3: fixes x :: "32 word" shows "\x \ y; x \ y; y < 2 ^ 32 - 1\ \ x \ y - 1" by (rule le_step_down_word_2, assumption+) lemma shiftr_1: "(x::word32) >> 1 = 0 \ x < 2" by transfer (simp add: take_bit_drop_bit drop_bit_Suc) lemma has_zero_byte: "~~ (((((v::word32) && 0x7f7f7f7f) + 0x7f7f7f7f) || v) || 0x7f7f7f7f) \ 0 \ v && 0xff000000 = 0 \ v && 0xff0000 = 0 \ v && 0xff00 = 0 \ v && 0xff = 0" by word_bitwise auto lemma mask_step_down_32: \\x. mask x = b\ if \b && 1 = 1\ and \\x. x < 32 \ mask x = b >> 1\ for b :: \32word\ proof - from \b && 1 = 1\ have \odd b\ by (auto simp add: mod_2_eq_odd and_one_eq) then have \b mod 2 = 1\ using odd_iff_mod_2_eq_one by blast from \\x. x < 32 \ mask x = b >> 1\ obtain x where \x < 32\ \mask x = b >> 1\ by blast then have \mask x = b div 2\ using shiftr1_is_div_2 [of b] by simp with \b mod 2 = 1\ have \2 * mask x + 1 = 2 * (b div 2) + b mod 2\ by (simp only:) also have \\ = b\ by (simp add: mult_div_mod_eq) finally have \2 * mask x + 1 = b\ . moreover have \mask (Suc x) = 2 * mask x + (1 :: 'a::len word)\ by (simp add: mask_Suc_rec) ultimately show ?thesis by auto qed lemma unat_of_int_32: "\i \ 0; i \2 ^ 31\ \ (unat ((of_int i)::sword32)) = nat i" unfolding unat_eq_nat_uint apply (subst eq_nat_nat_iff) apply (auto simp add: take_bit_int_eq_self) done lemmas word_ctz_not_minus_1_32 = word_ctz_not_minus_1[where 'a=32, simplified] (* Helper for packing then unpacking a 64-bit variable. *) lemma cast_chunk_assemble_id_64[simp]: "(((ucast ((ucast (x::64 word))::32 word))::64 word) || (((ucast ((ucast (x >> 32))::32 word))::64 word) << 32)) = x" by (simp add:cast_chunk_assemble_id) (* Another variant of packing and unpacking a 64-bit variable. *) lemma cast_chunk_assemble_id_64'[simp]: "(((ucast ((scast (x::64 word))::32 word))::64 word) || (((ucast ((scast (x >> 32))::32 word))::64 word) << 32)) = x" by (simp add:cast_chunk_scast_assemble_id) (* Specialisations of down_cast_same for adding to local simpsets. *) lemma cast_down_u64: "(scast::64 word \ 32 word) = (ucast::64 word \ 32 word)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done lemma cast_down_s64: "(scast::64 sword \ 32 word) = (ucast::64 sword \ 32 word)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done lemma word32_and_max_simp: \x AND 0xFFFFFFFF = x\ for x :: \32 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) end diff --git a/thys/Word_Lib/Word_64.thy b/thys/Word_Lib/Word_64.thy --- a/thys/Word_Lib/Word_64.thy +++ b/thys/Word_Lib/Word_64.thy @@ -1,332 +1,335 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 64" theory Word_64 imports Word_Lemmas Word_8 Word_16 Word_Syntax + Rsplit begin type_synonym word64 = "64 word" lemma len64: "len_of (x :: 64 itself) = 64" by simp +type_synonym sword64 = "64 sword" + type_synonym machine_word_len = 64 type_synonym machine_word = "machine_word_len word" definition word_bits :: nat where "word_bits = LENGTH(machine_word_len)" text \The following two are numerals so they can be used as nats and words.\ definition word_size_bits :: "'a :: numeral" where "word_size_bits = 3" definition word_size :: "'a :: numeral" where "word_size = 8" lemma word_bits_conv[code]: "word_bits = 64" unfolding word_bits_def by simp lemma word_size_word_size_bits: "(word_size::nat) = 2 ^ word_size_bits" unfolding word_size_def word_size_bits_def by simp lemma word_bits_word_size_conv: "word_bits = word_size * 8" unfolding word_bits_def word_size_def by simp lemma ucast_8_64_inj: "inj (ucast :: 8 word \ 64 word)" by (rule down_ucast_inj) (clarsimp simp: is_down_def target_size source_size) lemma upto_2_helper: "{0..<2 :: 64 word} = {0, 1}" by (safe; simp) unat_arith lemmas upper_bits_unset_is_l2p_64 = upper_bits_unset_is_l2p [where 'a=64, folded word_bits_def] lemmas le_2p_upper_bits_64 = le_2p_upper_bits [where 'a=64, folded word_bits_def] lemmas le2p_bits_unset_64 = le2p_bits_unset[where 'a=64, folded word_bits_def] lemma word_bits_len_of: "len_of TYPE (64) = word_bits" by (simp add: word_bits_conv) lemmas unat_power_lower64' = unat_power_lower[where 'a=64] lemmas unat_power_lower64 [simp] = unat_power_lower64'[unfolded word_bits_len_of] lemmas word64_less_sub_le' = word_less_sub_le[where 'a = 64] lemmas word64_less_sub_le[simp] = word64_less_sub_le' [folded word_bits_def] lemma word_bits_size: "size (w::word64) = word_bits" by (simp add: word_bits_def word_size) lemmas word64_power_less_1' = word_power_less_1[where 'a = 64] lemmas word64_power_less_1[simp] = word64_power_less_1'[folded word_bits_def] lemma of_nat64_0: "\of_nat n = (0::word64); n < 2 ^ word_bits\ \ n = 0" by (erule of_nat_0, simp add: word_bits_def) lemma unat_mask_2_less_4: "unat (p && mask 2 :: word64) < 4" apply (rule unat_less_helper) apply (rule order_le_less_trans, rule word_and_le1) apply (simp add: mask_eq) done lemmas unat_of_nat64' = unat_of_nat_eq[where 'a=64] lemmas unat_of_nat64 = unat_of_nat64'[unfolded word_bits_len_of] lemmas word_power_nonzero_64 = word_power_nonzero [where 'a=64, folded word_bits_def] lemmas unat_mult_simple = iffD1 [OF unat_mult_lem [where 'a = 64, unfolded word_bits_len_of]] lemmas div_power_helper_64 = div_power_helper [where 'a=64, folded word_bits_def] lemma n_less_word_bits: "(n < word_bits) = (n < 64)" by (simp add: word_bits_def) lemmas of_nat_less_pow_64 = of_nat_power [where 'a=64, folded word_bits_def] lemma lt_word_bits_lt_pow: "sz < word_bits \ sz < 2 ^ word_bits" by (simp add: word_bits_conv) lemma unat_less_word_bits: fixes y :: word64 shows "x < unat y \ x < 2 ^ word_bits" unfolding word_bits_def by (rule order_less_trans [OF _ unat_lt2p]) lemmas unat_mask_word64' = unat_mask[where 'a=64] lemmas unat_mask_word64 = unat_mask_word64'[folded word_bits_def] lemma unat_less_2p_word_bits: "unat (x :: 64 word) < 2 ^ word_bits" apply (simp only: word_bits_def) apply (rule unat_lt2p) done lemma Suc_unat_mask_div: "Suc (unat (mask sz div word_size::word64)) = 2 ^ (min sz word_bits - 3)" apply (case_tac "sz < word_bits") apply (case_tac "3 \ sz") apply (clarsimp simp: word_size_def word_bits_def min_def mask_eq) apply (drule (2) Suc_div_unat_helper [where 'a=64 and sz=sz and us=3, simplified, symmetric]) apply (simp add: not_le word_size_def word_bits_def) apply (case_tac sz, simp add: unat_word_ariths) apply (case_tac nat, simp add: unat_word_ariths unat_mask_word64 min_def word_bits_def) apply (case_tac nata, simp add: unat_word_ariths unat_mask_word64 word_bits_def) apply simp apply (simp add: unat_word_ariths unat_mask_word64 min_def word_bits_def word_size_def) done lemmas word64_minus_one_le' = word_minus_one_le[where 'a=64] lemmas word64_minus_one_le = word64_minus_one_le'[simplified] lemma ucast_not_helper: fixes a::word8 assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::word64)" proof assume "ucast a = (0xFF::word64)" also have "(0xFF::word64) = ucast (0xFF::word8)" by simp finally show False using a apply - apply (drule up_ucast_inj, simp) apply simp done qed lemma less_4_cases: "(x::word64) < 4 \ x=0 \ x=1 \ x=2 \ x=3" apply clarsimp apply (drule word_less_cases, erule disjE, simp, simp)+ done lemma if_then_1_else_0: "((if P then 1 else 0) = (0 :: word64)) = (\ P)" by simp lemma if_then_0_else_1: "((if P then 0 else 1) = (0 :: word64)) = (P)" by simp lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 lemma ucast_le_ucast_8_64: "(ucast x \ (ucast y :: word64)) = (x \ (y :: word8))" by (simp add: ucast_le_ucast) lemma in_16_range: "0 \ S \ r \ (\x. r + x * (16 :: word64)) ` S" "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: word64. r + x * 16) ` S" by (clarsimp simp: image_def elim!: bexI[rotated])+ lemma eq_2_64_0: "(2 ^ 64 :: word64) = 0" by simp lemma x_less_2_0_1: fixes x :: word64 shows "x < 2 \ x = 0 \ x = 1" by (rule x_less_2_0_1') auto lemmas mask_64_max_word = max_word_mask [symmetric, where 'a=64, simplified] lemma of_nat64_n_less_equal_power_2: "n < 64 \ ((of_nat n)::64 word) < 2 ^ n" by (rule of_nat_n_less_equal_power_2, clarsimp simp: word_size) lemma word_rsplit_0: "word_rsplit (0 :: word64) = [0, 0, 0, 0, 0, 0, 0, 0 :: word8]" by (simp add: word_rsplit_def bin_rsplit_def) lemma unat_ucast_10_64 : fixes x :: "10 word" shows "unat (ucast x :: word64) = unat x" by transfer simp lemma bool_mask [simp]: fixes x :: word64 shows "(0 < x && 1) = (x && 1 = 1)" by (rule bool_mask') auto lemma word64_bounds: "- (2 ^ (size (x :: word64) - 1)) = (-9223372036854775808 :: int)" "((2 ^ (size (x :: word64) - 1)) - 1) = (9223372036854775807 :: int)" "- (2 ^ (size (y :: 64 signed word) - 1)) = (-9223372036854775808 :: int)" "((2 ^ (size (y :: 64 signed word) - 1)) - 1) = (9223372036854775807 :: int)" by (simp_all add: word_size) lemma word_ge_min:"sint (x::64 word) \ -9223372036854775808" by (metis sint_ge word64_bounds(1) word_size) lemmas signed_arith_ineq_checks_to_eq_word64' = signed_arith_ineq_checks_to_eq[where 'a=64] signed_arith_ineq_checks_to_eq[where 'a="64 signed"] lemmas signed_arith_ineq_checks_to_eq_word64 = signed_arith_ineq_checks_to_eq_word64' [unfolded word64_bounds] lemmas signed_mult_eq_checks64_to_64' = signed_mult_eq_checks_double_size[where 'a=64 and 'b=64] signed_mult_eq_checks_double_size[where 'a="64 signed" and 'b=64] lemmas signed_mult_eq_checks64_to_64 = signed_mult_eq_checks64_to_64'[simplified] lemmas sdiv_word64_max' = sdiv_word_max [where 'a=64] sdiv_word_max [where 'a="64 signed"] lemmas sdiv_word64_max = sdiv_word64_max'[simplified word_size, simplified] lemmas sdiv_word64_min' = sdiv_word_min [where 'a=64] sdiv_word_min [where 'a="64 signed"] lemmas sdiv_word64_min = sdiv_word64_min' [simplified word_size, simplified] lemmas sint64_of_int_eq' = sint_of_int_eq [where 'a=64] lemmas sint64_of_int_eq = sint64_of_int_eq' [simplified] lemma ucast_of_nats [simp]: "(ucast (of_nat x :: word64) :: sword64) = (of_nat x)" "(ucast (of_nat x :: word64) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word64) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word8) :: sword8) = (of_nat x)" by (simp_all add: of_nat_take_bit take_bit_word_eq_self) lemmas signed_shift_guard_simpler_64' = power_strict_increasing_iff[where b="2 :: nat" and y=31] lemmas signed_shift_guard_simpler_64 = signed_shift_guard_simpler_64'[simplified] lemma word64_31_less: "31 < len_of TYPE (64 signed)" "31 > (0 :: nat)" "31 < len_of TYPE (64)" "31 > (0 :: nat)" by auto lemmas signed_shift_guard_to_word_64 = signed_shift_guard_to_word[OF word64_31_less(1-2)] signed_shift_guard_to_word[OF word64_31_less(3-4)] lemma le_step_down_word_3: fixes x :: "64 word" shows "\x \ y; x \ y; y < 2 ^ 64 - 1\ \ x \ y - 1" by (rule le_step_down_word_2, assumption+) lemma shiftr_1: "(x::word64) >> 1 = 0 \ x < 2" by transfer (simp add: take_bit_drop_bit drop_bit_Suc) lemma mask_step_down_64: \\x. mask x = b\ if \b && 1 = 1\ and \\x. x < 64 \ mask x = b >> 1\ for b :: \64word\ proof - from \b && 1 = 1\ have \odd b\ by (auto simp add: mod_2_eq_odd and_one_eq) then have \b mod 2 = 1\ using odd_iff_mod_2_eq_one by blast from \\x. x < 64 \ mask x = b >> 1\ obtain x where \x < 64\ \mask x = b >> 1\ by blast then have \mask x = b div 2\ using shiftr1_is_div_2 [of b] by simp with \b mod 2 = 1\ have \2 * mask x + 1 = 2 * (b div 2) + b mod 2\ by (simp only:) also have \\ = b\ by (simp add: mult_div_mod_eq) finally have \2 * mask x + 1 = b\ . moreover have \mask (Suc x) = 2 * mask x + (1 :: 'a::len word)\ by (simp add: mask_Suc_rec) ultimately show ?thesis by auto qed lemma unat_of_int_64: "\i \ 0; i \ 2 ^ 63\ \ (unat ((of_int i)::sword64)) = nat i" unfolding unat_eq_nat_uint apply (subst eq_nat_nat_iff) apply (simp_all add: take_bit_int_eq_self) done lemmas word_ctz_not_minus_1_64 = word_ctz_not_minus_1[where 'a=64, simplified] (* Helper for packing then unpacking a 64-bit variable. *) lemma cast_chunk_assemble_id_64[simp]: "(((ucast ((ucast (x::64 word))::32 word))::64 word) || (((ucast ((ucast (x >> 32))::32 word))::64 word) << 32)) = x" by (simp add:cast_chunk_assemble_id) (* Another variant of packing and unpacking a 64-bit variable. *) lemma cast_chunk_assemble_id_64'[simp]: "(((ucast ((scast (x::64 word))::32 word))::64 word) || (((ucast ((scast (x >> 32))::32 word))::64 word) << 32)) = x" by (simp add:cast_chunk_scast_assemble_id) (* Specialisations of down_cast_same for adding to local simpsets. *) lemma cast_down_u64: "(scast::64 word \ 32 word) = (ucast::64 word \ 32 word)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done lemma cast_down_s64: "(scast::64 sword \ 32 word) = (ucast::64 sword \ 32 word)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done lemma word64_and_max_simp: \x AND 0xFFFFFFFFFFFFFFFF = x\ for x :: \64 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) end diff --git a/thys/Word_Lib/Word_8.thy b/thys/Word_Lib/Word_8.thy --- a/thys/Word_Lib/Word_8.thy +++ b/thys/Word_Lib/Word_8.thy @@ -1,106 +1,108 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 8" theory Word_8 imports Word_Lemmas Word_Syntax begin type_synonym word8 = "8 word" lemma len8: "len_of (x :: 8 itself) = 8" by simp +type_synonym sword8 = "8 sword" + lemma word8_and_max_simp: \x AND 0xFF = x\ for x :: \8 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma enum_word8_eq: \enum = [0 :: 8 word, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255]\ (is \?lhs = ?rhs\) proof - have \map unat ?lhs = [0..<256]\ by (simp add: enum_word_def comp_def take_bit_nat_eq_self map_idem_upt_eq) also have \\ = map unat ?rhs\ by (simp add: upt_zero_numeral_unfold) finally show ?thesis using unat_inj by (rule map_injective) qed lemma set_enum_word8_def: "(set enum :: 8 word set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by (simp add: enum_word8_eq) lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: \8 word\ shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done end diff --git a/thys/Word_Lib/Word_EqI.thy b/thys/Word_Lib/Word_EqI.thy --- a/thys/Word_Lib/Word_EqI.thy +++ b/thys/Word_Lib/Word_EqI.thy @@ -1,150 +1,71 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Solving Word Equalities" theory Word_EqI imports - Word_Lib - Aligned More_Word_Operations "HOL-Eisbach.Eisbach_Tools" begin text \ Some word equalities can be solved by considering the problem bitwise for all @{prop "n < LENGTH('a::len)"}, which is different to running @{text word_bitwise} and expanding into an explicit list of bits. \ -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 test_bit_over: - "n \ size (x::'a::len word) \ (x !! n) = False" - by transfer auto - -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 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 le_mask_high_bits: - "w \ mask n \ (\i \ {n ..< size w}. \ w !! i)" - for w :: \'a::len word\ - by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) - -lemma neg_mask_le_high_bits: - "NOT(mask n) \ w \ (\i \ {n ..< size w}. w !! i)" - for w :: \'a::len word\ - by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_test_bit) - -lemma test_bit_conj_lt: - "(x !! m \ m < LENGTH('a)) = x !! m" for x :: "'a :: len word" - using test_bit_bin by blast - -lemma neg_test_bit: - "(NOT x) !! n = (\ 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) - named_theorems word_eqI_simps lemmas [word_eqI_simps] = word_ops_nth_size word_size word_or_zero neg_mask_test_bit nth_ucast is_aligned_nth nth_w2p nth_shiftl nth_shiftr less_2p_is_upper_bits_unset le_mask_high_bits neg_mask_le_high_bits bang_eq neg_test_bit is_up is_down + sign_extend_bitwise_if' -lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size] - -lemmas word_eqI_rule = word_eqI[rule_format] +lemmas word_eqI_rule = word_eqI [rule_format] lemma test_bit_lenD: "x !! n \ n < LENGTH('a) \ x !! n" for x :: "'a :: len word" by (fastforce dest: test_bit_size simp: word_size) method word_eqI uses simp simp_del split split_del cong flip = ((* reduce conclusion to test_bit: *) rule word_eqI_rule, (* make sure we're in clarsimp normal form: *) (clarsimp simp: simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* turn x < 2^n assumptions into mask equations: *) ((drule less_mask_eq)+)?, (* expand and distribute test_bit everywhere: *) (clarsimp simp: word_eqI_simps simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* add any additional word size constraints to new indices: *) ((drule test_bit_lenD)+)?, (* try to make progress (can't use +, would loop): *) (clarsimp simp: word_eqI_simps simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* helps sometimes, rarely: *) (simp add: simp test_bit_conj_lt del: simp_del flip: flip split: split split del: split_del cong: cong)?) method word_eqI_solve uses simp simp_del split split_del cong flip = solves \word_eqI simp: simp simp_del: simp_del split: split split_del: split_del cong: cong simp flip: flip; (fastforce dest: test_bit_size simp: word_eqI_simps simp flip: flip simp: simp simp del: simp_del split: split split del: split_del cong: cong)?\ 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,4357 +1,4358 @@ (* * 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 - Next_and_Prev Signed_Division_Word - Enumeration_Word - Word_EqI - Rsplit + Signed_Words More_Divides + More_Word Even_More_List More_Misc Most_significant_bit + Enumeration_Word + Aligned + Word_EqI begin text \Lemmas about words\ lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done 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 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) (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by transfer (simp add: bit_take_bit_iff ac_simps) 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) - by word_eqI_solve + 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 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 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: word_unat_power del: of_nat_power) + by (simp add: word_unat_power of_nat_diff del: of_nat_power) 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 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 of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) 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 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 of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed 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 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 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 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]) - apply word_eqI - apply (erule_tac x="na + n" in allE) - apply fastforce + 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 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 dest: bit_imp_le_length) 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]) - apply word_eqI - apply (erule_tac x="na - n" in allE) - apply auto + 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 ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x AND mask (len_of TYPE ('a))" - by word_eqI + 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('b \ 'a) x) = 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 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 up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (unfold scast_eq) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp 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 nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: rev_nth) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) 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 cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done 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 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 shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done lemma if_and_helper: "(If x v v') AND v'' = If x (v AND v'') (v' AND v'')" by (rule if_distrib) 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 lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] 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 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\ apply (simp add: and_not_mask) apply (case_tac "n 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 eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a AND mask b = (mask (min a b) :: 'a::len word)" - by word_eqI + 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\ - by safe word_eqI_solve + apply transfer + apply (simp add: bit_eq_iff) + apply (auto simp add: bit_simps ac_simps) + done lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n, simp) apply (case_tac n; simp) done lemma word_and_1_shiftl: "x AND (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" - by word_eqI_solve + 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\ - by word_eqI_solve + apply (rule bit_word_eqI; transfer) + apply (auto simp add: bit_simps not_le ac_simps) + done lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) 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 (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_eq of_nat_nat[symmetric] take_bit_eq_mod) 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 range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp 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 nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done 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 i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done 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 If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp 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 word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] 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 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 of_bool_nth: "of_bool (x !! v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) 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 shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" - by word_eqI_solve + 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 dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) 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 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 (simp add: word_unat.norm_eq_iff [symmetric]) done lemma mask_AND_NOT_mask: "(w AND NOT (mask n)) AND mask n = 0" for w :: \'a::len word\ by word_eqI lemma AND_NOT_mask_plus_AND_mask_eq: "(w AND NOT (mask n)) + (w AND mask n) = w" for w :: \'a::len word\ by (subst word_plus_and_or_coroll; word_eqI_solve) 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 (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x AND mask n) !! m)" by (simp add: word_size test_bit_conj_lt) also have "\ = ((y AND mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size test_bit_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x AND NOT (mask n)) !! m)" by (simp add: neg_mask_test_bit test_bit_conj_lt) also have "\ = ((y AND NOT (mask n)) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_test_bit test_bit_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) 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_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 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 and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x AND (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply word_eqI done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x AND 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) 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 (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done 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 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 sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" apply (cases \LENGTH('a)\) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) apply (metis le_less_Suc_eq test_bit_bin test_bit_word_eq) 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 drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp 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 not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp 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) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* 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 (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp 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 sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc 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 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) (* Signed word arithmetic overflow constraints. *) 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)+ 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: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (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 simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" - apply (auto simp: sdiv_int_def sgn_if) + apply (auto simp: signed_divide_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" - by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) + by (auto simp: signed_divide_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) - apply (clarsimp simp: sdiv_int_def) + apply (clarsimp simp: signed_divide_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) - apply (clarsimp simp: sdiv_int_def sgn_if) + apply (clarsimp simp: signed_divide_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" - apply (clarsimp simp: sdiv_int_def) + apply (clarsimp simp: signed_divide_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less algebra_split_simps) apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" - apply (unfold sdiv_int_def) + apply (unfold signed_divide_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (auto simp add: sgn_if not_less) apply (metis le_less le_less_trans neg_equal_0_iff_equal neg_less_iff_less not_le pos_imp_zdiv_neg_iff) apply (metis add.inverse_neutral div_int_pos_iff le_less neg_le_iff_le order_trans) apply (metis div_minus_right le_less_trans neg_imp_zdiv_neg_iff neg_less_0_iff_less not_le) using div_int_pos_iff apply fastforce apply (metis abs_0_eq abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) - apply (clarsimp simp: sdiv_word_def sdiv_int_def) - apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) + apply (clarsimp simp: sdiv_word_def signed_divide_int_def) + apply (clarsimp simp: sdiv_word_def signed_divide_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" - by (clarsimp simp: sdiv_int_def) + by (clarsimp simp: signed_divide_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" - by (clarsimp simp: sdiv_int_def) + by (clarsimp simp: signed_divide_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" - apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) + apply (auto simp: sdiv_word_def signed_divide_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" - apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) + apply (auto simp: sdiv_word_def signed_divide_int_def sgn_if) done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done 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)}" 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) 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: sdiv_int_def sgn_if word_size sint_int_min) + 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) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num 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] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where v="numeral x" for x] sdiv_word_def[where v=0] sdiv_word_def[where v=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where w="numeral y" for y] word_sdiv_numerals_lhs[where w=0] word_sdiv_numerals_lhs[where w=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" - apply (clarsimp simp: smod_int_def sdiv_int_def) + apply (clarsimp simp: signed_modulo_int_def signed_divide_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" - by (clarsimp simp: smod_int_def) + by (clarsimp simp: signed_modulo_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (cases \a \ - (2 ^ (LENGTH('a) - 1)) \ b \ - 1\) - apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def + apply (clarsimp simp: smod_word_def sdiv_word_def signed_modulo_int_def simp flip: wi_hom_sub wi_hom_mult) - apply (clarsimp simp: smod_word_def smod_int_def) + apply (clarsimp simp: smod_word_def signed_modulo_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where v="numeral x" for x] smod_word_def[where v=0] smod_word_def[where v=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where w="numeral y" for y] word_smod_numerals_lhs[where w=0] word_smod_numerals_lhs[where w=1] 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 nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" by transfer (auto simp add: bit_simps) lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (fact ucast_nat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) 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_mod_power_int min_def) apply (rule distrib [symmetric]) 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_bl 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 smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done 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 (subst sint_eq_uint) apply (simp add: msb_word_eq) apply transfer apply (simp add: bit_take_bit_iff) apply transfer apply simp 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 (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) 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 add: Power.of_nat_power) 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 add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" by transfer (simp add: bit_simps) lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a AND mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply transfer apply (cases \LENGTH('a)\) apply (simp_all add: take_bit_Suc_from_most bit_simps) apply (simp_all add: bit_simps disjunctive_add) done -lemma word_rsplit_upt: - "\ size x = LENGTH('a :: len) * n; n \ 0 \ - \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" - apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") - apply (rule nth_equalityI, simp) - apply (intro allI word_eqI impI) - apply (simp add: test_bit_rsplit_alt word_size) - apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) - apply (simp add: length_word_rsplit_exp_size) - apply (metis mult.commute given_quot_alt word_size word_size_gt_0) - done - lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma neg_mask_add_mask: "((x:: 'a :: len word) AND NOT (mask n)) + (2 ^ n - 1) = x OR mask n" unfolding mask_2pm1[symmetric] by (subst word_plus_and_or_coroll; word_eqI_solve) 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 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 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 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 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 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 shiftl_word_eq bit_push_bit_iff) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith 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 (* 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 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 (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) 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 (auto simp add: Parity.bit_eq_iff bit_xor_iff 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 scast_nop1 [simp]: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2 [simp]: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop = scast_nop1 scast_nop2 scast_id 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 (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff bit_mask_iff) 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 bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by transfer (simp add: drop_bit_Suc) 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 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 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 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 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 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 \ x !! 0" using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = 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 simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(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':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_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 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 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 word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) 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_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 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 complement_nth_w2p: shows "n' < LENGTH('a) \ (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 td_gal_lt) 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_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply transfer apply simp done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done lemma word_add_no_overflow:"(x::'a::len word) < max_word \ 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] (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed 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 shiftr_eq_0) 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) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (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_sub_int) 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 (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp 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: "(max_word :: 'a::len word) = mask LENGTH('a)" unfolding mask_eq_decr_exp by simp 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 word_eqI_solve 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 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 word_eqI_solve 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 word_eqI_solve lemma ucast_and_neg_mask: "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" by word_eqI_solve lemma ucast_and_mask: "ucast (x AND mask n) = ucast x AND mask n" by word_eqI_solve lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" by word_eqI lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by (simp add: is_down ucast_ucast_a) 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 word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce 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 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 from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" by (simp add: bit_eq_iff test_bit_word_eq) 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') 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: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed 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 (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) 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) by (meson less_irrefl not_less zle2p) 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: \bin_cat a n b \ 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: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc 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 (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ 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 (* 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 unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) 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_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) 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 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 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 - word_eqI thus ?thesis by simp qed lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" apply (induct xs) apply simp apply (simp add: of_bl_Cons mask_Suc) apply (rule conjI; clarsimp) apply (erule word_plus_mono_right) apply (rule is_aligned_no_overflow_mask) apply (rule is_aligned_triv) apply (simp add: word_le_nat_alt) apply (subst unat_add_lem') apply (rule is_aligned_mask_offset_unat) apply (rule is_aligned_triv) apply (simp add: mask_eq_decr_exp) apply simp done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) 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 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 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 power_2_mult_step_le: "\n' \ n; 2 ^ n' * k' < 2 ^ n * k\ \ 2 ^ n' * (k' + 1) \ 2 ^ n * (k::nat)" apply (cases "n'=n", simp) apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7)) apply (drule (1) le_neq_trans) apply clarsimp apply (subgoal_tac "\m. n = n' + m") prefer 2 apply (simp add: le_Suc_ex) apply (clarsimp simp: power_add) by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) 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 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 word_eqI_solve 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" by (clarsimp simp: word_eqI_simps) 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" by (clarsimp simp: le_mask_high_bits word_size nth_ucast) 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 + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp 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 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 max_word_not_less[simp]: "\ max_word < x" by (simp add: not_less) 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 fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) 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: unat_eq_of_nat) + 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" 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 ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve 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 pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by simp lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat word_unat.Rep_inverse) 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 neg_mask_combine: "NOT(mask a) AND NOT(mask b) = NOT(mask (max a b) :: 'a::len word)" by (auto simp: word_ops_nth_size word_size intro!: word_eqI) 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 (metis neg_mask_combine) 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 add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' AND NOT(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis le_word_or2 neg_mask_add_mask and.right_idem) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (simp only: is_aligned_add_or flip: neg_mask_in_mask_range) by (metis less_mask_eq mask_subsume) lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_eq_decr_exp) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_eq_decr_exp) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) apply auto apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_eq_decr_exp) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply (intro range_subsetI) apply (rule is_aligned_no_wrap' [OF al xsz]) apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps) lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p AND NOT(mask n') \ p'; p' AND NOT(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x AND mask n >> m" in spec) apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1) done 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) 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 small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using small_powers_of_2 by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2 not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower word_clz_max word_of_nat_less wsst_TYs(3)) lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using small_powers_of_2 uint_nat apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less word_size) by (simp add: uint_nat) 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" using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt by (metis (no_types, lifting)) 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" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) 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_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma of_bl_length2: "length xs + c < LENGTH('a) \ of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) 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: Word_Lemmas.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) Word_Lemmas.of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) 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 from_to_bool_last_bit: "from_bool (to_bool (x AND 1)) = x AND 1" by (metis from_bool_to_bool_iff word_and_1) 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 Word_Lemmas.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: word_size nth_w2p split: if_split) 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 (max_word :: '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) 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 (metis cast_simps(23) scast_scast_id(2)) 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_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: bit_iff_odd) 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 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)" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower) lemma sint_ctz: "LENGTH('a) > 2 \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") apply (rule conjI) apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) by (rule small_powers_of_2, simp) 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) 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 unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_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 add: mask_eq_decr_exp NOT_eq flip: mul_not_mask_eq_neg_shiftl) 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 (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" by (metis assms sbintrunc_bintrunc ucast_eq word_ubin.eq_norm) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. w !! i = 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))" by word_eqI_solve 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 of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) AND NOT(mask n) = a" by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1)) lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (subst word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule is_aligned_AND_less_0) apply (simp add: is_aligned_mask) apply (rule order_less_le_trans) apply (rule of_bl_length2) apply simp apply (simp add: two_power_increasing) apply simp apply (rule nth_equalityI) apply (simp only: len_bin_to_bl) apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric]) apply (simp add: nth_shiftr nth_shiftl shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric]) apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl rev_nth) apply arith done (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) -declare of_nat_power [simp del] - 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 end diff --git a/thys/Word_Lib/Word_Lib.thy b/thys/Word_Lib/Word_Lib.thy deleted file mode 100644 --- a/thys/Word_Lib/Word_Lib.thy +++ /dev/null @@ -1,97 +0,0 @@ -(* - * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) - * - * SPDX-License-Identifier: BSD-2-Clause - *) - -section "Additional Word Operations" - -theory Word_Lib - imports - "HOL-Library.Signed_Division" - More_Word - Signed_Words - Traditional_Infix_Syntax - Reversed_Bit_Lists -begin - -(* Haskellish names/syntax *) -notation (input) - test_bit ("testBit") - -lemma scast_scast_id [simp]: - "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" - "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" - by (auto simp: is_up scast_up_scast_id) - -lemma scast_ucast_id [simp]: - "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" - by (metis down_cast_same is_down len_signed order_refl scast_scast_id(1)) - -lemma ucast_scast_id [simp]: - "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" - by (metis scast_scast_id(2) scast_ucast_id) - -lemma scast_of_nat [simp]: - "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" - by transfer (simp add: take_bit_signed_take_bit) - -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 word_sint_1 [simp]: - "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" - by (fact signed_1) - -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 scast_eq_scast_id [simp]: - "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" - by (metis ucast_scast_id) - -lemma ucast_eq_ucast_id [simp]: - "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" - by (metis scast_ucast_id) - -lemma scast_ucast_norm [simp]: - "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" - "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" - by (metis scast_ucast_id ucast_scast_id)+ - -(* FIXME: these should eventually be moved to HOL/Word. *) -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 - -(* shortcut for some specific lengths *) -lemma word_fixed_sint_1[simp]: - "sint (1::8 word) = 1" - "sint (1::16 word) = 1" - "sint (1::32 word) = 1" - "sint (1::64 word) = 1" - by (auto simp: sint_word_ariths) - -lemmas sdiv_int_def = signed_divide_int_def -lemmas smod_int_def = signed_modulo_int_def - -end diff --git a/thys/Word_Lib/Word_Lib_Sumo.thy b/thys/Word_Lib/Word_Lib_Sumo.thy --- a/thys/Word_Lib/Word_Lib_Sumo.thy +++ b/thys/Word_Lib/Word_Lib_Sumo.thy @@ -1,98 +1,114 @@ section \Ancient comprehensive Word Library\ theory Word_Lib_Sumo imports "HOL-Library.Word" Aligned Ancient_Numeral Bit_Comprehension Bits_Int Bitwise_Signed Bitwise Enumeration_Word Generic_set_bit Hex_Words Least_significant_bit More_Arithmetic More_Divides More_Sublist Even_More_List More_Misc Strict_part_mono Legacy_Aliasses Most_significant_bit Next_and_Prev Norm_Words Reversed_Bit_Lists Rsplit Signed_Words Traditional_Infix_Syntax Typedef_Morphisms Type_Syntax Word_EqI - Word_Lib Word_Lemmas Word_8 Word_16 Word_32 Word_64 Word_Syntax Signed_Division_Word More_Word_Operations begin declare signed_take_bit_Suc [simp] lemmas bshiftr1_def = bshiftr1_eq lemmas is_down_def = is_down_eq lemmas is_up_def = is_up_eq lemmas mask_def = mask_eq_decr_exp lemmas scast_def = scast_eq lemmas shiftl1_def = shiftl1_eq lemmas shiftr1_def = shiftr1_eq lemmas sshiftr1_def = sshiftr1_eq lemmas sshiftr_def = sshiftr_eq_funpow_sshiftr1 lemmas to_bl_def = to_bl_eq lemmas ucast_def = ucast_eq lemmas unat_def = unat_eq_nat_uint lemmas word_cat_def = word_cat_eq lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl lemmas word_rotl_def = word_rotl_eq lemmas word_rotr_def = word_rotr_eq lemmas word_sle_def = word_sle_eq lemmas word_sless_def = word_sless_eq lemmas uint_0 = uint_nonnegative lemmas uint_lt = uint_bounded lemmas uint_mod_same = uint_idem lemmas of_nth_def = word_set_bits_def lemmas of_nat_word_eq_iff = word_of_nat_eq_iff lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff lemmas of_int_word_eq_iff = word_of_int_eq_iff lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff lemmas word_next_def = word_next_unfold lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat lemma shiftl_transfer [transfer_rule]: includes lifting_syntax shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" by (unfold shiftl_eq_push_bit) transfer_prover lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp word32_and_max_simp word64_and_max_simp lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemmas and_bang = word_and_nth +lemmas sdiv_int_def = signed_divide_int_def +lemmas smod_int_def = signed_modulo_int_def + +(* shortcut for some specific lengths *) +lemma word_fixed_sint_1[simp]: + "sint (1::8 word) = 1" + "sint (1::16 word) = 1" + "sint (1::32 word) = 1" + "sint (1::64 word) = 1" + by (auto simp: sint_word_ariths) + +declare of_nat_diff [simp] + +(* Haskellish names/syntax *) +notation (input) + test_bit ("testBit") + end