Theory More_Word

(*
 * Copyright Data61, CSIRO (ABN 41 687 119 230)
 *
 * SPDX-License-Identifier: BSD-2-Clause
Proofs tidied by LCP, 2024-09
 *)

section ‹Lemmas on words›

theory More_Word
  imports
    "HOL-Library.Word" More_Arithmetic More_Divides More_Bit_Ring
begin

declare bit_1_iff[simp]

context
  includes bit_operations_syntax
begin

― ‹problem posed by TPHOLs referee:
      criterion for overflow of addition of signed integers›

lemma sofl_test:
  sint x + sint y = sint (x + y) 
    drop_bit (size x - 1) ((x + y XOR x) AND (x + y XOR y)) = 0
  for x y :: 'a::len word
proof -
  obtain n where n: LENGTH('a) = Suc n
    by (cases LENGTH('a)) simp_all
  have *: sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y)  sint x + sint y  - (2 ^ n)
    signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n  2 ^ n > sint x + sint y
    using signed_take_bit_int_greater_eq [of sint x + sint y n] signed_take_bit_int_less_eq [of n sint x + sint y]
    by (auto intro: ccontr)
  have sint x + sint y = sint (x + y) 
    (sint (x + y) < 0  sint x < 0)  (sint (x + y) < 0  sint y < 0)
    by (smt (verit) One_nat_def add_diff_cancel_left' signed_take_bit_int_eq_self sint_greater_eq sint_lt sint_word_ariths(1,2))
  then show ?thesis
    unfolding One_nat_def word_size drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff
    by (simp add: bit_last_iff)
qed

lemma unat_power_lower [simp]:
  "unat ((2::'a::len word) ^ n) = 2 ^ n" if "n < LENGTH('a::len)"
  using that by transfer simp

lemma unat_p2: "n < LENGTH('a :: len)  unat (2 ^ n :: 'a word) = 2 ^ n"
  by (fact unat_power_lower)

lemma word_div_lt_eq_0:
  "x < y  x div y = 0" for x :: "'a :: len word"
  by (fact div_word_less)

lemma word_div_eq_1_iff: "n div m = 1  n  m  unat n < 2 * unat (m :: 'a :: len word)"
  by (metis One_nat_def nat_div_eq_Suc_0_iff of_nat_unat ucast_id unat_1 unat_div word_le_nat_alt)

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 (fact word_pow_0)

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: 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: 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
  by (metis and_zero_eq less_mask_eq word_bw_lcs(1))

lemma and_mask_eq_iff_le_mask:
  w AND mask n = w  w  mask n
  for w :: 'a::len word
  by (smt (verit) and.idem mask_eq_iff word_and_le1 word_le_def)

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"
  by (fact not_eq_complement)

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
  using le_m1_iff_lt by blast

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)"
  by (metis mod_less order_less_trans unat_of_nat word_less_nat_alt xlt)

lemma word_and_max_word:
  fixes a:: "'a::len word"
  shows "x = - 1  a AND x = a"
  by simp

lemma word_and_full_mask_simp [simp]:
  x AND mask LENGTH('a) = x for x :: 'a::len word
  by (simp add: bit_eq_iff bit_simps)

lemma of_int_uint [simp]: "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
  by (fact disjunctive_add2)

corollary word_plus_and_or_coroll2:
  "(x AND w) + (x AND NOT w) = x"
  for x w :: 'a::len word
  by (fact and_plus_not_and)

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"
  by (meson le_trans up_ucast_inj)

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"
    using nv unat_mono yv by force
  have ux: "unat x < 2 ^ m"
    using mv unat_mono xv by fastforce
  have "unat x < 2 ^ (LENGTH('a :: len) - n)"
    by (metis exp_eq_zero_iff not_less0 linorder_not_le unat_mono unat_power_lower unsigned_0 xv')
  then have *: "unat x * 2 ^ n < 2 ^ LENGTH('a)"
    by (simp add: nat_mult_power_less_eq)
  show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' *
    apply (simp add: word_less_nat_alt unat_word_ariths)
    by (metis less_imp_diff_less mn mod_nat_add nat_add_offset_less unat_power_lower unsigned_less)
qed

lemma word_less_power_trans:
  fixes n :: "'a :: len word"
  assumes "n < 2 ^ (m - k)" "k  m" "m < len_of TYPE ('a)"
  shows "2 ^ k * n < 2 ^ m"
proof -
  have "2 ^ k * unat n < 2 ^ LENGTH('a)"
  proof -
    have "(1::nat) < 2"
      by simp
    moreover
    have "m - k < len_of (TYPE('a)::'a itself)"
      by (simp add: assms less_imp_diff_less)
    with assms have "unat n < 2 ^ (m - k)"
      by (metis (no_types) unat_power_lower word_less_iff_unsigned)
    ultimately show ?thesis
      by (meson assms order.strict_trans nat_less_power_trans power_strict_increasing)
  qed
  then show ?thesis
    using assms nat_less_power_trans
    by (simp add: word_less_nat_alt unat_word_ariths)
qed

lemma  word_less_power_trans2:
  fixes n :: "'a::len word"
  shows "n < 2 ^ (m - k); k  m; m < LENGTH('a)  n * 2 ^ k < 2 ^ m"
  by (subst field_simps, rule word_less_power_trans)

lemma Suc_unat_diff_1:
  fixes x :: "'a :: len word"
  assumes lt: "1  x"
  shows "Suc (unat (x - 1)) = unat x"
  by (metis Suc_diff_1 linorder_not_less lt unat_gt_0 unat_minus_one word_less_1)

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 "y  x" "0 < y"
shows "(x - y) div y = x div y - 1"
  using assms  by (simp add: word_div_def div_pos_geq uint_minus_simple_alt uint_sub_lem word_less_def)

lemma word_mult_less_mono1:
  fixes i :: "'a :: len word"
  assumes "i < j" and "0 < k"
    and "unat j * unat k < 2 ^ len_of TYPE ('a)"
  shows  "i * k < j * k"
  by (simp add: assms div_lt_mult word_div_mult)

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"
    by (simp add: le_div_geq unat_div usv)

  also have " = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv
    by (simp add: unat_minus_one)
  also have " = (2 ^ us - 1 + 2 ^ us * 2 ^ q) div 2 ^ us"
    by (simp add: power_add qv)
  also have " = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)"
    by (metis (no_types) not_less_zero div_mult_self2 take_bit_nat_less_exp)
  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 "i  j" and "unat j + unat k < 2 ^ len_of TYPE ('a)"
  shows  "i + k  j + k"
  using assms no_olen_add_nat word_plus_mono_left by fastforce

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 (metis add.commute no_olen_add_nat word_plus_mono_right)

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)"
  using assms word_add_le_dest word_add_le_mono1 by blast

lemma word_add_less_mono1:
  fixes i :: "'a :: len word"
  assumes "i < j" and "unat j + unat k < 2 ^ len_of TYPE ('a)"
  shows  "i + k < j + k"
  using assms no_olen_add_nat not_less_iff_gr_or_eq olen_add_eqv word_l_diffs(2) by fastforce

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)"
  using assms word_add_less_dest word_add_less_mono1 by blast

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"  "0 < k"
  and "unat j * unat k < 2 ^ len_of TYPE ('a)"
  shows  "i * k  j * k"
  by (simp add: assms div_le_mult word_div_mult)

lemma word_mult_le_iff:
  fixes i :: "'a :: len word"
  assumes "0 < k"
  and     "unat i * unat k < 2 ^ len_of TYPE ('a)"
  and     "unat j * unat k < 2 ^ len_of TYPE ('a)"
  shows  "(i * k  j * k) = (i  j)"
  by (metis assms div_le_mult nle_le word_div_mult)


lemma word_diff_less:
  fixes n :: "'a :: len word"
  shows "0 < n; 0 < m; n  m  m - n < m"
  by (metis linorder_not_le sub_wrap word_greater_zero_iff)

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"
  using word_gt_0 word_not_simps(1) by blast

lemma word_power_less_1 [simp]:
  "sz < LENGTH('a::len)  (2::'a word) ^ sz - 1 < 2 ^ sz"
  using power_2_ge_iff by blast

lemma word_sub_1_le:
  "x  0  x - 1  (x :: 'a :: len word)"
  by (simp add: word_le_sub1 word_sub_le)

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 unsigned_of_nat)
      (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: unsigned_of_nat 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: 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 assms
    by (metis le_unat_uoi mult.commute nat_le_linear nat_less_power_trans not_less unat_power_lower)

  have "unat a div 2 ^ n * 2 ^ n  unat a"
    using assms
    by (metis Abs_fnat_hom_0 mod_mult_self2_is_0 unat_power_lower word_arith_nat_mod)

  then have "a div 2 ^ n * 2 ^ n < a" using assms
    by (metis add_cancel_right_right le_less word_div_mult_le word_mod_div_equality)

  also from xk le have "  of_nat k * 2 ^ n" by (simp add: field_simps)
  finally show ?thesis using sz kv
    by (smt (verit) div_mult_le kn order_le_less_trans unat_div unsigned_less word_mult_less_dest)
qed

lemma mask_out_sub_mask:
  "(x AND NOT (mask n)) = x - (x AND (mask n))"
  for x :: 'a::len word
  by (fact and_not_eq_minus_and)

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 (auto simp: and_not_eq_minus_and)

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)" "y < LENGTH('a)"
  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"
  using inc_le linorder_not_le by blast

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:
  "x  - 1  (y < x + 1) = (y < x  y = x)" for x :: "'a::len word"
  by (meson max_word_wrap plus_one_helper plus_one_helper2 word_le_less_eq)

lemma word_Suc_leq:
  fixes k:: "'a::len word" shows "k  - 1  x < k + 1  x  k"
  using less_x_plus_1 word_le_less_eq by auto

lemma word_Suc_le:
   fixes k:: "'a::len word" shows "x  - 1  x + 1  k  x < k"
  by (meson not_less word_Suc_leq)

lemma word_lessThan_Suc_atMost:
  {..< k + 1} = {..k} if k  - 1 for k :: 'a::len word
  using that by (simp add: lessThan_def atMost_def word_Suc_leq)

lemma word_atLeastLessThan_Suc_atLeastAtMost:
  {l ..< u + 1} = {l..u} if u  - 1 for l :: 'a::len word
  using that by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost)

lemma word_atLeastAtMost_Suc_greaterThanAtMost:
  {m<..u} = {m + 1..u} if m  - 1 for m :: 'a::len word
  using that by (simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le)

lemma word_atLeastLessThan_Suc_atLeastAtMost_union:
  fixes l:: "'a::len word"
  assumes "m  - 1" and "l  m" and "m  u"
  shows "{l..m}  {m+1..u} = {l..u}"
  by (metis assms ivl_disj_un_two(8) word_atLeastAtMost_Suc_greaterThanAtMost)

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 (verit) suc_n le_less_trans lessI nat_less_le nat_mult_less_cancel_disj p2_gt_0
          power_Suc power_Suc unat_power_lower word_less_nat_alt x)

lemma power_overflow:
  "n  LENGTH('a)  2 ^ n = (0 :: 'a::len word)"
  by simp

lemmas extra_sle_sless_unfolds [simp] =
    word_sle_eq[where a=0 and b=1]
    word_sle_eq[where a=0 and b="numeral n"]
    word_sle_eq[where a=1 and b=0]
    word_sle_eq[where a=1 and b="numeral n"]
    word_sle_eq[where a="numeral n" and b=0]
    word_sle_eq[where a="numeral n" and b=1]
    word_sless_alt[where a=0 and b=1]
    word_sless_alt[where a=0 and b="numeral n"]
    word_sless_alt[where a=1 and b=0]
    word_sless_alt[where a=1 and b="numeral n"]
    word_sless_alt[where a="numeral n" and b=0]
    word_sless_alt[where a="numeral n" and b=1]
  for n

lemma word_sint_1:
  "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)"
  by (fact signed_1)

lemma ucast_of_nat:
  "is_down (ucast :: 'a :: len word  'b :: len word)
     ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)"
  by transfer simp

lemma scast_1':
  "(scast (1::'a::len word) :: 'b::len word) =
   (word_of_int (signed_take_bit (LENGTH('a::len) - Suc 0) (1::int)))"
  by transfer simp

lemma scast_1:
  "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)"
  by (fact signed_1)

lemma unat_minus_one_word:
  "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1"
  by (simp add: mask_eq_exp_minus_1 unsigned_minus_1_eq_mask)

lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x]
lemmas word_diff_ls' = word_diff_ls'' [simplified]

lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x]
lemmas word_l_diffs = word_l_diffs' [simplified]

lemma two_power_increasing:
  " n  m; m < LENGTH('a)   (2 :: 'a :: len word) ^ n  2 ^ m"
  by (simp add: word_le_nat_alt)

lemma word_leq_le_minus_one:
  " x  y; x  0   x - 1 < (y :: 'a :: len word)"
  by (meson le_m1_iff_lt linorder_not_less word_greater_zero_iff)

lemma neg_mask_combine:
  "NOT(mask a) AND NOT(mask b) = NOT(mask (max a b) :: 'a::len word)"
  by (rule bit_word_eqI) (auto simp: bit_simps)

lemma neg_mask_twice:
  "x AND NOT(mask n) AND NOT(mask m) = x AND NOT(mask (max n m))"
  for x :: 'a::len word
  by (rule bit_word_eqI) (auto simp: bit_simps)

lemma multiple_mask_trivia:
  "n  m  (x AND NOT(mask n)) + (x AND mask n AND NOT(mask m)) = x AND NOT(mask m)"
  for x :: 'a::len word
  by (metis (no_types, lifting) add.commute add_diff_eq and.assoc and_not_eq_minus_and and_plus_not_and mask_twice min_def)

lemma word_of_nat_less: "n < unat x  of_nat n < x"
  by (metis le_unat_uoi nat_less_le word_less_nat_alt)

lemma unat_mask:
  "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1"
  by (metis mask_eq_exp_minus_1 min.commute unat_mask_eq)

lemma mask_over_length:
  "LENGTH('a)  n  mask n = (-1::'a::len word)"
  by (simp add: mask_eq_decr_exp)

lemma Suc_2p_unat_mask:
  "n < LENGTH('a)  Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)"
  by (simp add: unat_mask)

lemma sint_of_nat_ge_zero:
  "x < 2 ^ (LENGTH('a) - 1)  sint (of_nat x :: 'a :: len word)  0"
  by (simp add: bit_iff_odd signed_of_nat)

lemma int_eq_sint:
  assumes "x < 2 ^ (LENGTH('a) - 1)"
  shows "sint (of_nat x :: 'a :: len word) = int x"
proof -
  have "int x < 2 ^ (len_of (TYPE('a)::'a itself) - 1)"
    by (metis assms of_nat_less_iff of_nat_numeral of_nat_power)
  then show ?thesis
    by (smt (verit) One_nat_def id_apply of_int_eq_id of_nat_0_le_iff signed_of_nat signed_take_bit_int_eq_self)
qed

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 (simp add: int_eq_sint less_imp_diff_less)

lemma word_le_not_less:
  fixes b :: "'a::len word"
  shows "b  a  ¬ a < b"
  by fastforce

lemma less_is_non_zero_p1:
  fixes a :: "'a :: len word"
  shows "a < k  a + 1  0"
  using linorder_not_le max_word_wrap by auto

lemma unat_add_lem':
  fixes y :: "'a::len word"
  shows "(unat x + unat y < 2 ^ LENGTH('a))  (unat (x + y) = unat x + unat y)"
  using unat_add_lem by blast

lemma word_less_two_pow_divI:
  " (x :: 'a::len word) < 2 ^ (n - m); m  n; n < LENGTH('a)   x < 2 ^ n div 2 ^ m"
  by (simp add: word_less_nat_alt power_minus_is_div unat_div)

lemma word_less_two_pow_divD:
  fixes x :: "'a::len word"
  assumes "x < 2 ^ n div 2 ^ m"
  shows "n  m  (x < 2 ^ (n - m))"
proof -
  have f2: "unat x < unat ((2::'a word) ^ n div 2 ^ m)"
    using assms by (simp add: word_less_nat_alt)
  then have f3: "0 < unat ((2::'a word) ^ n div 2 ^ m)"
    using order_le_less_trans by blast
  have f4: "n < LENGTH('a)"
    by (metis assms div_0 possible_bit_def possible_bit_word word_zero_le [THEN leD])
  then have "2 ^ n div 2 ^ m = unat ((2::'a word) ^ n div 2 ^ m)"
    by (metis div_by_0 exp_eq_zero_iff f3 linorder_not_le unat_div unat_gt_0 unat_power_lower)
  then show ?thesis
    by (metis assms f2 power_minus_is_div two_pow_div_gt_le unat_div word_arith_nat_div word_unat_power)
qed

lemma of_nat_less_two_pow_div_set:
  assumes "n < LENGTH('a)"
  shows  "{x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}"
proof -
  have "k<2 ^ n div 2 ^ m. w = word_of_nat k"
    if "w < 2 ^ n div 2 ^ m" for w :: "'a word"
    using that assms
    by (metis less_imp_diff_less power_minus_is_div unat_less_power unat_of_nat_len word_less_two_pow_divD word_nchotomy)
  moreover have "(word_of_nat k::'a word) < 2 ^ n div 2 ^ m"
    if "k < 2 ^ n div 2 ^ m" for k
    using that assms   
    by (metis order_le_less_trans two_pow_div_gt_le unat_div unat_power_lower word_of_nat_less)
  ultimately show ?thesis
    by (auto simp: word_of_nat_less)
qed

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 (simp add: ucast_less)
  by (metis (mono_tags, opaque_lifting) UNIV_I Word.of_nat_unat image_eqI unat_eq_of_nat unat_less_power unat_lt2p)

lemma word_power_less_diff:
  fixes q :: "'a::len word"
  assumes  "2 ^ n * q < 2 ^ m" and "q < 2 ^ (LENGTH('a) - n)"
    shows "q < 2 ^ (m - n)"
proof (cases "m  LENGTH('a)  n  LENGTH('a)  n=0")
  case True
  then show ?thesis
    using assms
    by (elim context_disjE; simp add: power_overflow)
next
  case False
  have "2 ^ n * unat q < 2 ^ m"
    by (metis assms p2_gt_0 unat_eq_of_nat unat_less_power unat_lt2p unat_mult_power_lem word_gt_a_gt_0)
  with assms have "unat q < 2 ^ (m - n)"
    using nat_power_less_diff by blast
  then show ?thesis
    using False word_less_nat_alt by fastforce
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)"
  using add_diff_cancel_left' word_le_minus_mono by fastforce

lemma word_not_le:
  "(¬ x  (y :: 'a :: len word)) = (y < x)"
  by (fact not_le)

lemma word_subset_less:
  fixes s :: "'a :: len word"
  assumes "{x..x + r - 1}  {y..y + s - 1}"
      and xy: "x  x + r - 1" "y  y + s - 1"
      and "s  0"
    shows "r  s"
proof -
  obtain "x  y + (s - 1)" "y  x" "x + (r - 1)  y + (s - 1)"
    using assms by (auto simp flip: add_diff_eq)
  then have "r - 1  s - 1"
    by (metis add_diff_eq xy olen_add_eqv word_sub_mono2)
  then show ?thesis
    using s  0 word_le_minus_cancel word_le_sub1 by auto
qed

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"
  by (simp add: word_le_nat_alt)

lemma two_power_eq:
  "n < LENGTH('a); m < LENGTH('a)
    ((2::'a::len word) ^ n = 2 ^ m) = (n = m)"
  by (metis nle_le power_le_mono)

lemma unat_less_helper:
  "x < of_nat n  unat x < n"
  by (metis not_less_iff_gr_or_eq word_less_nat_alt word_of_nat_less)

lemma nat_uint_less_helper:
  "nat (uint y) = z  x < y  nat (uint x) < z"
  using nat_less_eq_zless uint_lt_0 word_less_iff_unsigned by blast

lemma of_nat_0:
  "of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)  n = 0"
  by (auto simp: word_of_nat_eq_0_iff)

lemma of_nat_inj:
  "x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a) 
   (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)"
  by (metis unat_of_nat_len)

lemma div_to_mult_word_lt:
  " (x :: 'a :: len word)  y div z   x * z  y"
  by (cases "z = 0") (simp_all add: div_le_mult word_neq_0_conv)

lemma ucast_ucast_mask:
  "(ucast :: 'a :: len word  'b :: len word) (ucast x) = x AND mask (len_of TYPE ('a))"
  by (metis Word.of_int_uint and_mask_bintr unsigned_ucast_eq)

lemma ucast_ucast_len:
  " x < 2 ^ LENGTH('b)   ucast (ucast x::'b::len word) = (x::'a::len word)"
  by (simp add: less_mask_eq ucast_ucast_mask)

lemma ucast_ucast_id:
  "LENGTH('a) < LENGTH('b)  ucast (ucast (x::'a::len word)::'b::len word) = x"
  using is_up less_or_eq_imp_le ucast_up_ucast_id by blast

lemma unat_ucast:
  "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))"
  by (metis Word.of_nat_unat unat_of_nat)

lemma ucast_less_ucast:
  "LENGTH('a)  LENGTH('b) 
   (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)"
  by (metis Word.of_nat_unat is_up not_less_iff_gr_or_eq ucast_up_ucast_id word_of_nat_less)

― ‹This weaker version was previously called @{text ucast_less_ucast}. We retain it to
    support existing proofs.›
lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order]

lemma unat_Suc2:
  fixes n :: "'a :: len word"
  shows "n  -1  unat (n + 1) = Suc (unat n)"
  by (metis add.commute max_word_wrap unatSuc)

lemma word_div_1:
  "(n :: 'a :: len word) div 1 = n"
  by (fact div_by_1)

lemma word_minus_one_le:
  "-1  (x :: 'a :: len word) = (x = -1)"
  by (fact word_order.extremum_unique)

lemma up_scast_inj:
  " scast x = (scast y :: 'b :: len word); size x  LENGTH('b)   x = y"
  by (metis is_up scast_up_scast_id word_size)

lemma up_scast_inj_eq:
  "LENGTH('a)  len_of TYPE ('b) 
  (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))"
  by (metis is_up scast_up_scast_id)

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':
  "x + v  x + w; x  x + v  v  w" for x :: "'a::len word"
  by (meson olen_add_eqv order_refl word_add_increasing word_sub_mono2)

lemma unat_eq_1:
  unat x = Suc 0  x = 1
  by (auto intro!: unsigned_word_eqI [where ?'a = nat])

lemma word_unat_Rep_inject1:
  unat x = unat 1  x = 1
  by (simp add: unat_eq_1)

lemma and_not_mask_twice:
  "(w AND NOT (mask n)) AND NOT (mask m) = w AND NOT (mask (max m n))"
  for w :: 'a::len word
  by (rule bit_word_eqI) (auto simp: bit_simps)

lemma word_less_cases:
  "x < y  x = y - 1  x < y - (1 ::'a::len word)"
  by (meson order_le_imp_less_or_eq word_le_minus_one_leq)

lemma mask_and_mask:
  "mask a AND mask b = (mask (min a b) :: 'a::len word)"
  by (simp flip: take_bit_eq_mask ac_simps)

lemma mask_eq_0_eq_x:
  "(x AND w = 0) = (x AND NOT w = x)"
  for x w :: 'a::len word
  by (simp add: and_not_eq_minus_and)

lemma mask_eq_x_eq_0:
  "(x AND w = x) = (x AND NOT w = 0)"
  for x w :: 'a::len word
  by (metis and_not_eq_minus_and eq_iff_diff_eq_0)

lemma compl_of_1: "NOT 1 = (-2 :: 'a :: len word)"
  by (fact not_one_eq)

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 (metis word_bw_comms(1) word_plus_and_or_coroll2)

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)"
  by (metis Word.of_nat_unat of_nat_inverse)

lemma word_le_make_less:
  fixes x :: "'a :: len word"
  shows "y  -1  (x  y) = (x < (y + 1))"
  by (simp add: word_Suc_leq)

lemmas finite_word = finite [where 'a="'a::len word"]

lemma word_to_1_set:
  "{0 ..< (1 :: 'a :: len word)} = {0}"
  by fastforce

lemma word_leq_minus_one_le:
  fixes x :: "'a::len word"
  shows "y  0; x  y - 1   x < y"
  using le_m1_iff_lt word_neq_0_conv by blast

lemma word_count_from_top:
  "n  0  {0 ..< n :: 'a :: len word} = {0 ..< n - 1}  {n - 1}"
  using word_leq_minus_one_le word_less_cases by force

lemma word_minus_one_le_leq:
  " x - 1 < y   x  (y :: 'a :: len word)"
  using diff_add_cancel inc_le by force

lemma word_must_wrap:
  " x  n - 1; n  x   n = (0 :: 'a :: len word)"
  using dual_order.trans sub_wrap word_less_1 by blast

lemma range_subset_card:
  " {a :: 'a :: len word .. b}  {c .. d}; b  a   d  c  d - c  b - a"
  using word_sub_le word_sub_mono by fastforce

lemma less_1_simp:
  "n - 1 < m = (n  (m :: 'a :: len word)  n  0)"
  by unat_arith

lemma word_power_mod_div:
  fixes x :: "'a::len word"
  shows " n < LENGTH('a); m < LENGTH('a)
   x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)"
  by (metis drop_bit_eq_div drop_bit_take_bit take_bit_eq_mod)

lemma word_range_minus_1':
  fixes a :: "'a :: len word"
  shows "a  0  {a-1<..b} = {a..b}"
  by (simp add: word_atLeastAtMost_Suc_greaterThanAtMost)

lemma word_range_minus_1:
  fixes a :: "'a :: len word"
  shows "b  0  {a..b - 1} = {a..<b}"
  by (auto simp: word_le_minus_one_leq word_leq_minus_one_le)

lemma ucast_nat_def:
  "of_nat (unat x) = (ucast :: 'a :: len word  'b :: len word) x"
  by transfer simp

lemma overflow_plus_one_self:
  "(1 + p  p) = (p = (-1 :: 'a :: len word))"
  by (metis add.commute order_less_irrefl word_Suc_le word_order.extremum)

lemma plus_1_less:
  "(x + 1  (x :: 'a :: len word)) = (x = -1)"
  using word_Suc_leq by blast

lemma pos_mult_pos_ge:
  "[|x > (0::int); n>=0 |] ==> n * x >= n*1"
  by (simp add: mult_le_cancel_left1)

lemma word_plus_strict_mono_right:
  fixes x :: "'a :: len word"
  shows "y < z; x  x + z  x + y < x + z"
  by unat_arith

lemma word_div_mult:
  "0 < c  a < b * c  a div c < b" for a b c :: "'a::len word"
  by (metis antisym_conv3 div_lt_mult leD order.asym word_div_mult_le)

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"
  by (simp add: word_less_power_trans2 word_of_nat_less)

lemma word_1_le_power:
  "n < LENGTH('a)  (1 :: 'a :: len word)  2 ^ n"
  by (metis bot_nat_0.extremum power_0 two_power_increasing)

lemma unat_1_0:
  "1  (x::'a::len word) = (0 < unat x)"
  by (auto simp: 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"
  by (metis One_nat_def less_2_cases of_nat_numeral unat_less_helper unsigned_0 unsigned_1 unsigned_word_eqI)

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"
  by (simp add: word_of_nat_less)

lemma of_nat_n_less_equal_power_2:
  "n < LENGTH('a::len)  ((of_nat n)::'a word) < 2 ^ n"
  by (simp add: More_Word.of_nat_power)

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 (metis and_mask_less' eqm 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))"
  by (simp add: unat_of_nat_len word_less_nat_alt xlt ylt)

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)"
  by (metis unat_of_nat_len word_less_eq_iff_unsigned)

lemma mask_AND_NOT_mask:
  "(w AND NOT (mask n)) AND mask n = 0"
  for w :: 'a::len word
  by (simp add: mask_eq_0_eq_x)

lemma AND_NOT_mask_plus_AND_mask_eq:
  "(w AND NOT (mask n)) + (w AND mask n) = w"
  for w :: 'a::len word
  by (simp add: add.commute word_plus_and_or_coroll2)

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"
  using m1 m2 split_word_eq_on_mask by blast

lemma neq_0_no_wrap:
  fixes x :: "'a :: len word"
  shows " x  x + y; x  0   x + y  0"
  by clarsimp

lemma unatSuc2:
  fixes n :: "'a :: len word"
  shows "n + 1  0  unat (n + 1) = Suc (unat n)"
  by (simp add: add.commute unatSuc)

lemma word_of_nat_le:
  "n  unat x  of_nat n  x"
  by (simp add: le_unat_uoi word_le_nat_alt)

lemma word_unat_less_le:
  "a  of_nat b  unat a  b"
  by (metis eq_iff le_cases le_unat_uoi word_of_nat_le)

lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)"
  by (simp add: mask_eq_decr_exp)

lemma bool_mask':
  fixes x :: "'a :: len word"
  shows "2 < LENGTH('a)  (0 < x AND 1) = (x AND 1 = 1)"
  by (simp add: and_one_eq mod_2_eq_odd)

lemma ucast_ucast_add:
  fixes x :: "'a :: len word"
  fixes y :: "'b :: len word"
  shows "LENGTH('b)  LENGTH('a)  ucast (ucast x + y) = x + ucast y"
  by transfer (smt (verit, ccfv_threshold) min_def take_bit_add take_bit_take_bit)

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"
  using max_word_wrap by fastforce

lemma word_sub_plus_one_nonzero:
  fixes n :: "'a :: len word"
  shows "n'  n; n'  0  (n - n') + 1  0"
  by (metis diff_add_cancel word_plus_one_nonzero word_sub_le)

lemma word_le_minus_mono_right:
  fixes x :: "'a :: len word"
  shows " z  y; y  x; z  x   x - y  x - z"
  using range_subset_card by auto

lemma word_0_sle_from_less:
  0 ≤s x if x < 2 ^ (LENGTH('a) - 1) for x :: 'a::len word
  using that
  by (metis p2_gt_0 signed_0 sint_of_nat_ge_zero unat_eq_of_nat unat_less_power unat_lt2p word_gt_a_gt_0 word_sle_eq)

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)"
  by (metis Word.of_nat_unat assms(1) of_nat_diff unat_sub word_less_eq_iff_unsigned)

lemma word_1_0:
  "a + (1::('a::len) word)  b; a < of_nat x  a < b"
  by (metis word_Suc_le word_order.extremum_strict)

lemma unat_of_nat_less: " a < b; unat b = c   a < of_nat c"
  by fastforce

lemma word_le_plus_1: " (y::('a::len) word) < y + n; a < n   y + a  y + a + 1"
  by unat_arith

lemma word_le_plus: "(a::('a::len) word) < a + b; c < b  a  a + c"
  by (metis order_less_imp_le word_random)

lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)"
  by (metis signed_word_eqI sint_n1)

lemma sint_0 [simp]: "(sint x = 0) = (x = 0)"
  by (fact signed_eq_0_iff)

(* It is not always that case that "sint 1 = 1", because of 1-bit word sizes.
 * This lemma produces the different cases. *)
lemma sint_1_cases:
  P if  len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0   P
      len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1   P
      len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1   P
proof (cases LENGTH('a) = 1)
  case True
  then have a = 0  a = 1
    by transfer auto
  with True that show ?thesis
    by auto
next
  case False
  with that show ?thesis
    by (simp add: less_le Suc_le_eq)
qed

lemma sint_int_min:
  "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))"
proof (cases LENGTH('a))
  case 0
  then show ?thesis
    by simp
next
  case Suc
  then show ?thesis
    by transfer (simp add: signed_take_bit_int_eq_self)
qed

lemma sint_int_max_plus_1:
  "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))"
proof (cases LENGTH('a))
  case 0
  then show ?thesis
    by simp
next
  case Suc
  then show ?thesis
    by transfer (simp add: signed_take_bit_eq_take_bit_shift take_bit_eq_mod word_of_int_2p)
qed

lemma uint_range': 0  uint x  uint x < 2 ^ LENGTH('a) for x :: 'a::len word
  by simp

lemma sint_of_int_eq:
  " - (2 ^ (LENGTH('a) - 1))  x; x < 2 ^ (LENGTH('a) - 1)   sint (of_int x :: ('a::len) word) = x"
  by (simp add: signed_take_bit_int_eq_self signed_of_int)

lemma of_int_sint:
  "of_int (sint a) = a"
  by simp

lemma sint_ucast_eq_uint:
    " ¬ is_down (ucast :: ('a::len word  'b::len word)) 
             sint ((ucast :: ('a::len word  'b::len word)) x) = uint x"
  by transfer (simp add: signed_take_bit_take_bit)

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)"
  by (metis Word.of_nat_unat of_nat_eq_iff uint_up_ucast)

lemma ucast_mono:
  " (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) 
    ucast x < ((ucast y) :: 'a :: len word)"
  by (metis Word.of_nat_unat of_nat_mono_maybe p2_gt_0 unat_less_power unat_mono word_gt_a_gt_0)

lemma ucast_mono_le:
  "x  y; y < 2 ^ LENGTH('b)  (ucast (x :: 'a :: len word) :: 'b :: len word)  ucast y"
  by (metis order_class.order_eq_iff ucast_mono word_le_less_eq)

lemma ucast_mono_le':
  " unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x  y 
    ucast x  (ucast y :: 'b word)" for x y :: 'a::len word
  by (auto simp: word_less_nat_alt intro: ucast_mono_le)

lemma neg_mask_add_mask:
  "((x:: 'a :: len word) AND NOT (mask n)) + (2 ^ n - 1) = x OR mask n"
  by (simp add: mask_eq_decr_exp or_eq_and_not_plus)

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 (meson le_step_down_word)

lemma NOT_mask_AND_mask[simp]: "(w AND mask n) AND NOT (mask n) = 0"
  by (simp add: and.assoc)

lemma and_and_not[simp]: "(a AND b) AND NOT b = 0"
  for a b :: 'a::len word
  using AND_twice mask_eq_x_eq_0 by blast

lemma ex_mask_1[simp]: "(x. mask x = (1 :: 'a::len word))"
  using mask_1 by blast

lemma not_switch: "NOT a = x  a = NOT x"
  by auto

lemma test_bit_eq_iff: "bit u = bit v  u = v"
  for u v :: "'a::len word"
  by (auto intro: bit_eqI simp add: fun_eq_iff)

lemma test_bit_size: "bit w n  n < size w"
  for w :: "'a::len word"
  by transfer simp

lemma word_eq_iff: "x = y  (n<LENGTH('a). bit x n = bit y n)"
  for x y :: "'a::len word"
  using bit_word_eqI by blast

lemma word_eqI: "(n. n < size u  bit u n = bit v n)  u = v"
  for u :: "'a::len word"
  by (simp add: word_size word_eq_iff)

lemma word_eqD: "u = v  bit u x = bit v x"
  for u v :: "'a::len word"
  by simp

lemma test_bit_bin': "bit w n  n < size w  bit (uint w) n"
  by transfer (simp add: bit_take_bit_iff)

lemmas test_bit_bin = test_bit_bin' [unfolded word_size]

lemma word_test_bit_def: bit a = bit (uint a)
  using bit_uint_iff test_bit_bin by blast

lemmas test_bit_def' = word_test_bit_def [THEN fun_cong]

lemma word_test_bit_transfer [transfer_rule]:
  "(rel_fun pcr_word (rel_fun (=) (=)))
    (λx n. n < LENGTH('a)  bit x n) (bit :: 'a::len word  _)"
  by transfer_prover

lemma word_ops_nth_size:
  "n < size x 
    bit (x OR y) n = (bit x n | bit y n) 
    bit (x AND y) n = (bit x n  bit y n) 
    bit (x XOR y) n = (bit x n  bit y n) 
    bit (NOT x) n = (¬ bit x n)"
  for x :: "'a::len word"
  by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff)

lemma word_ao_nth:
  "bit (x OR y) n = (bit x n | bit y n) 
    bit (x AND y) n = (bit x n  bit y n)"
  for x :: "'a::len word"
  using bit_and_iff bit_or_iff by blast

lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]]

lemma nth_sint:
  fixes w :: "'a::len word"
  defines "l  LENGTH('a)"
  shows "bit (sint w) n = (if n < l - 1 then bit w n else bit w (l - 1))"
  unfolding sint_uint l_def
  by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def)

lemma test_bit_2p: "bit (word_of_int (2 ^ n)::'a::len word) m  m = n  m < LENGTH('a)"
  by transfer (auto simp: bit_exp_iff)

lemma nth_w2p: "bit ((2::'a::len word) ^ n) m  m = n  m < LENGTH('a::len)"
  by transfer (auto simp: bit_exp_iff)

lemma bang_is_le: "bit x m  2 ^ m  x"
  for x :: "'a::len word"
  by (metis bit_take_bit_iff less_irrefl linorder_le_less_linear take_bit_word_eq_self_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 nth_0: "¬ bit (0 :: 'a::len word) n"
  by simp

lemma nth_minus1: "bit (-1 :: 'a::len word) n  n < LENGTH('a)"
  by simp

lemma nth_ucast_weak:
  "bit (ucast w::'a::len word) n = (bit w n  n < LENGTH('a))"
  using bit_ucast_iff by blast

lemma nth_ucast:
  "bit (ucast (w::'a::len word)::'b::len word) n =
   (bit w n  n < min LENGTH('a) LENGTH('b))"
  by (metis bit_word_ucast_iff min_less_iff_conj nth_ucast_weak)

lemma nth_mask:
  bit (mask n :: 'a::len word) i  i < n  i < size (mask n :: 'a word)
  by (auto simp: word_size Word.bit_mask_iff)

lemma nth_slice: "bit (slice n w :: 'a::len word) m = (bit w (m + n)  m < LENGTH('a))"
  using bit_imp_le_length
  by (fastforce simp: bit_simps less_diff_conv dest: bit_imp_le_length)

― ‹keep quantifiers for use in simplification›
lemma test_bit_split':
  "word_split c = (a, b) 
    (n m.
      bit b n = (n < size b  bit c n) 
      bit a m = (m < size a  bit c (m + size b)))"
  by (auto simp: word_split_bin' bit_unsigned_iff word_size bit_drop_bit_eq ac_simps
           dest: bit_imp_le_length)

lemma test_bit_split:
  "word_split c = (a, b) 
    (n::nat. bit b n  n < size b  bit c n) 
    (m::nat. bit a m  m < size a  bit c (m + size b))"
  by (simp add: test_bit_split')

lemma test_bit_rcat:
  "sw = size (hd wl)  rc = word_rcat wl  bit rc n =
    (n < size rc  n div sw < size wl  bit ((rev wl) ! (n div sw)) (n mod sw))"
  for wl :: "'a::len word list"
  by (simp add: word_size word_rcat_def rev_map bit_horner_sum_uint_exp_iff bit_simps not_le)

lemmas test_bit_cong = arg_cong [where f = "bit", THEN fun_cong]

lemma max_test_bit: "bit (- 1::'a::len word) n  n < LENGTH('a)"
  by (fact nth_minus1)

lemma map_nth_0 [simp]: "map (bit (0::'a::len word)) xs = replicate (length xs) False"
  by (simp flip: map_replicate_const)

lemma word_and_1:
  "n AND 1 = (if bit n 0 then 1 else 0)" for n :: "_ word"
  by (rule bit_word_eqI) (auto simp: bit_and_iff bit_1_iff intro: gr0I)

lemma nth_w2p_same:
  "bit (2^n :: 'a :: len word) n = (n < LENGTH('a))"
  by (simp add: nth_w2p)

lemma word_leI:
  fixes u :: "'a::len word"
  assumes "n.  n < size u; bit u n   bit (v::'a::len word) n"
  shows "u  v"
proof (rule order_trans)
  show "u  u AND v"
    by (metis assms bit_and_iff word_eqI word_le_less_eq)
qed (simp add: word_and_le1)

lemma bang_eq:
  fixes x :: "'a::len word"
  shows "(x = y) = (n. bit x n = bit y n)"
  by (auto intro!: bit_eqI)

lemma neg_mask_test_bit:
  "bit (NOT(mask n) :: 'a :: len word) m = (n  m  m < LENGTH('a))"
  by (auto simp: bit_simps)

lemma upper_bits_unset_is_l2p:
  (n'  n. n' < LENGTH('a)  ¬ bit p n')  (p < 2 ^ n) (is ?P  ?Q)
    if n < LENGTH('a)
    for p :: "'a :: len word"
proof
  assume ?Q
  then show ?P
    by (meson bang_is_le le_less_trans not_le word_power_increasing)
next
  assume P: ?P
  have take_bit n p = p
  proof (rule bit_word_eqI)
    fix q
    assume q: 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: bit_simps)
    next
      case False
      then show ?thesis
        by (simp add: P q bit_take_bit_iff)
    qed
  qed
  with that show ?Q
    using take_bit_word_eq_self_iff [of n p] by auto
qed

lemma less_2p_is_upper_bits_unset:
  "p < 2 ^ n  n < LENGTH('a)  (n'  n. n' < LENGTH('a)  ¬ bit p n')" for p :: "'a :: len word"
  by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le)

lemma test_bit_over:
  "n  size (x::'a::len word)  (bit x n) = False"
  by transfer auto

lemma le_mask_high_bits:
  "w  mask n  (i  {n ..< size w}. ¬ bit w i)"
  for w :: 'a::len word
  by (metis atLeastLessThan_iff bit_take_bit_iff less_eq_mask_iff_take_bit_eq_self linorder_not_le nth_mask word_leI wsst_TYs(3))

lemma test_bit_conj_lt:
  "(bit x m  m < LENGTH('a)) = bit x m" for x :: "'a :: len word"
  using test_bit_bin by blast

lemma neg_test_bit:
  "bit (NOT x) n = (¬ bit x n  n < LENGTH('a))" for x :: "'a::len word"
  by (cases "n < LENGTH('a)") (auto simp: test_bit_over word_ops_nth_size word_size)

lemma nth_bounded:
  "bit (x :: 'a :: len word) n; x < 2 ^ m; m  len_of TYPE ('a)  n < m"
  by (meson less_2p_is_upper_bits_unset linorder_not_le test_bit_conj_lt)

lemma and_neq_0_is_nth:
  x AND y  0  bit x n if y = 2 ^ n for x y :: 'a::len word
  by (simp add: and_exp_eq_0_iff_not_bit that)

lemma nth_is_and_neq_0:
  "bit (x::'a::len word) n = (x AND 2 ^ n  0)"
  by (subst and_neq_0_is_nth; rule refl)

lemma max_word_not_less [simp]:
   "¬ - 1 < x" for x :: 'a::len word
  by (fact word_order.extremum_strict)

lemma bit_twiddle_min:
  "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y"
  by (rule bit_eqI) (auto simp: bit_simps)

lemma bit_twiddle_max:
  "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y"
  by (rule bit_eqI) (auto simp: bit_simps max_def)

lemma swap_with_xor:
  "(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y  z = b  y = a"
  by (auto intro: bit_word_eqI simp add: bit_simps)

lemma le_mask_imp_and_mask:
  "(x::'a::len word)  mask n  x AND mask n = x"
  by (metis and_mask_eq_iff_le_mask)

lemma or_not_mask_nop:
  "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n"
  by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3))

lemma mask_subsume:
  "n  m  ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)"
  by (rule bit_word_eqI) (auto simp: bit_simps word_size)

lemma and_mask_0_iff_le_mask:
  fixes w :: "'a::len word"
  shows "(w AND NOT(mask n) = 0) = (w  mask n)"
  by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask)

lemma mask_twice2:
  "n  m  ((x::'a::len word) AND mask m) AND mask n = x AND mask n"
  by (metis mask_twice min_def)

lemma uint_2_id:
  "LENGTH('a)  2  uint (2::('a::len) word) = 2"
  by simp

lemma div_of_0_id[simp]: "(0::('a::len) word) div n = 0"
  by (simp add: word_div_def)

lemma degenerate_word: "LENGTH('a) = 1  (x::('a::len) word) = 0  x = 1"
  by (metis One_nat_def less_irrefl_nat sint_1_cases)

lemma div_less_dividend_word:
  fixes x :: "('a::len) word"
  assumes "x  0" "n  1"
  shows "x div n < x"
proof (cases n = 0)
  case True
  then show ?thesis
    by (simp add: assms word_greater_zero_iff)
next
  case False
  then have "n > 1"
    by (metis assms(2) lt1_neq0 nless_le)
  then show ?thesis
    by (simp add: assms(1) unat_div unat_gt_0 word_less_nat_alt)
qed

lemma word_less_div:
  fixes x :: "('a::len) word"
    and y :: "('a::len) word"
  shows "x div y = 0  y = 0  x < y"
  by (metis div_greater_zero_iff linorder_not_le unat_div unat_gt_0 word_less_eq_iff_unsigned)

lemma not_degenerate_imp_2_neq_0: "LENGTH('a) > 1  (2::('a::len) word)  0"
  by (metis numerals(1) power_not_zero power_zero_numeral)

lemma word_overflow: "(x::('a::len) word) + 1 > x  x + 1 = 0"
  using plus_one_helper2 by auto

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))"
  by (metis even_plus_one_iff word_overflow_unat)

lemma odd_word_imp_even_next: "odd (unat (x::('a::len) word))  x + 1 = 0  even (unat (x + 1))"
  using even_plus_one_iff word_overflow_unat by force

lemma overflow_imp_lsb: "(x::('a::len) word) + 1 = 0  bit x 0"
  using even_plus_one_iff [of x] by (simp add: bit_0)

lemma odd_iff_lsb: "odd (unat (x::('a::len) word)) = bit x 0"
  by transfer (simp add: even_nat_iff bit_0)

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)"
  by (metis unat_of_nat)

lemma lsb_this_or_next: "¬ (bit ((x::('a::len) word) + 1) 0)  bit x 0"
  by (simp add: bit_0)

lemma mask_or_not_mask:
  "x AND mask n OR x AND NOT (mask n) = x"
  for x :: 'a::len word
  by (metis and.right_neutral bit.disj_cancel_right word_combine_masks)

lemma word_gr0_conv_Suc: "(m::'a::len word) > 0  n. m = n + 1"
  by (metis add.commute add_minus_cancel)

lemma revcast_down_us [OF refl]:
  "rc = revcast  source_size rc = target_size rc + n  rc w = ucast (signed_drop_bit n w)"
  for w :: "'a::len word"
  by (simp add: source_size_def target_size_def bit_word_eqI bit_simps ac_simps)

lemma revcast_down_ss [OF refl]:
  "rc = revcast  source_size rc = target_size rc + n  rc w = scast (signed_drop_bit n w)"
  for w :: "'a::len word"
  by (simp add: source_size_def target_size_def bit_word_eqI bit_simps ac_simps)

lemma revcast_down_uu [OF refl]:
  "rc = revcast  source_size rc = target_size rc + n  rc w = ucast (drop_bit n w)"
  for w :: "'a::len word"
  by (simp add: source_size_def target_size_def bit_word_eqI bit_simps ac_simps)

lemma revcast_down_su [OF refl]:
  "rc = revcast  source_size rc = target_size rc + n  rc w = scast (drop_bit n w)"
  for w :: "'a::len word"
  by (simp add: source_size_def target_size_def bit_word_eqI bit_simps ac_simps)

lemma cast_down_rev [OF refl]:
  "uc = ucast  source_size uc = target_size uc + n  uc w = revcast (push_bit n w)"
  for w :: "'a::len word"
  by (simp add: source_size_def target_size_def bit_word_eqI bit_simps)

lemma revcast_up [OF refl]:
  "rc = revcast  source_size rc + n = target_size rc 
    rc w = push_bit n (ucast w :: 'a::len word)"
  by (metis add_diff_cancel_left' le_add1 linorder_not_le revcast_def slice1_def wsst_TYs)

lemmas rc1 = revcast_up [THEN
  revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]]
lemmas rc2 = revcast_down_uu [THEN
  revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]]

lemma word_ops_nth:
  fixes x y :: 'a::len word
  shows
  word_or_nth:  "bit (x OR y) n = (bit x n  bit y n)" and
  word_and_nth: "bit (x AND y) n = (bit x n  bit y n)" and
  word_xor_nth: "bit (x XOR y) n = (bit x n  bit y n)"
  by (simp_all add: bit_simps)

lemma word_power_nonzero:
  " (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x  0 
   x * 2 ^ n  0"
  by (metis gr0I mult.commute not_less_eq p2_gt_0 power_0 word_less_1 word_power_less_diff zero_less_diff)

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"
  by (metis Suc_div_unat_helper add_diff_cancel_left' of_nat_Suc unat_div word_arith_nat_div word_unat_power)

lemma max_word_mask:
  "(- 1 :: 'a::len word) = mask LENGTH('a)"
  by (fact minus_1_eq_mask)

lemmas mask_len_max = max_word_mask[symmetric]

lemma mask_out_first_mask_some:
  " x AND NOT (mask n) = y; n  m   x AND NOT (mask m) = y AND NOT (mask m)"
  for x y :: 'a::len word
  by (rule bit_word_eqI) (auto simp: bit_simps word_size)

lemma mask_lower_twice:
  "n  m  (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)"
  for x :: 'a::len word
  by (rule bit_word_eqI) (auto simp: bit_simps word_size)

lemma mask_lower_twice2:
  "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))"
  for a :: 'a::len word
  by (simp add: and.assoc neg_mask_twice)

lemma ucast_and_neg_mask:
  "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)"
proof (rule bit_word_eqI)
  fix m 
  show "bit (ucast (x AND NOT (mask n))::'a word) m = bit ((ucast x::'a word) AND NOT (mask n)) m"
    by (smt (verit) bit_and_iff bit_word_ucast_iff neg_mask_test_bit)
qed

lemma ucast_and_mask:
  "ucast (x AND mask n) = ucast x AND mask n"
  by (metis take_bit_eq_mask unsigned_take_bit_eq)

lemma ucast_mask_drop:
  "LENGTH('a :: len)  n  (ucast (x AND mask n) :: 'a word) = ucast x"
  by (metis mask_twice2 ucast_and_mask word_and_full_mask_simp)

lemma mask_exceed:
  "n  LENGTH('a)  (x::'a::len word) AND NOT (mask n) = 0"
  by (rule bit_word_eqI) (simp add: bit_simps)

lemma word_add_no_overflow: "(x::'a::len word) < - 1  x < x + 1"
  using less_x_plus_1 order_less_le by blast

lemma lt_plus_1_le_word:
  fixes x :: "'a::len word"
  assumes bound: "n < unat (maxBound::'a word)"
  shows "x < 1 + of_nat n = (x  of_nat n)"
  by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less)

lemma unat_ucast_up_simp:
  fixes x :: "'a::len word"
  assumes "LENGTH('a)  LENGTH('b)"
  shows "unat (ucast x :: 'b::len word) = unat x"
  by (simp add: assms is_up unat_ucast_upcast)

lemma unat_ucast_less_no_overflow:
  "n < 2 ^ LENGTH('a); unat f < n  (f::('a::len) word) < of_nat n"
  by (simp add: unat_of_nat_len word_less_nat_alt)

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:
  fixes f :: "'a::len word" and b :: "'b :: len word"
  assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)"
  and upward_cast: "LENGTH('a) < LENGTH('b)"
  shows "(ucast f < b) = (unat f < unat b)"
proof -
  have LR: "ucast f < b  unat f < unat b"
    by (simp add: less_or_eq_imp_le unat_ucast_up_simp upward_cast word_less_nat_alt)
  have RL: "unat f < unat b  ucast f < b"
  proof-
    assume ineq: "unat f < unat b"
    have "ucast f < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)"
      by (metis ineq no_overflow ucast_nat_def unat_of_nat_len word_nchotomy word_of_nat_less)
    then show ?thesis
      using ineq word_of_nat_less by fastforce
  qed
  then show ?thesis by (simp add:RL LR iffI)
qed

lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2]

lemma minus_one_word:
  "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1"
  by simp

lemma le_2p_upper_bits:
  " (p::'a::len word)  2^n - 1; n < LENGTH('a)  
  n'n. n' < LENGTH('a)  ¬ bit p n'"
  by (subst upper_bits_unset_is_l2p; simp)

lemma le2p_bits_unset:
  "p  2 ^ n - 1  n'n. n' < LENGTH('a)  ¬ bit (p::'a::len word) n'"
  using upper_bits_unset_is_l2p [where p=p]
  by (cases "n < LENGTH('a)") auto

lemma complement_nth_w2p:
  shows "n' < LENGTH('a)  bit (NOT (2 ^ n :: 'a::len word)) n' = (n'  n)"
  by (fastforce simp: word_ops_nth_size word_size nth_w2p)

lemma word_unat_and_lt:
  "unat x < n  unat y < n  unat (x AND y) < n"
  by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt)

lemma word_unat_mask_lt:
  "m  size w  unat ((w::'a::len word) AND mask m) < 2 ^ m"
  by (rule word_unat_and_lt) (simp add: unat_mask word_size)

lemma word_sless_sint_le: "x <s y  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..<n}) = card {0..<n}"
  if "n  CARD('a::len word)"
proof -
  let ?of_nat = "of_nat :: nat  'a word"
  have "inj_on ?of_nat {i. i < CARD('a word)}"
    by (rule inj_onI) (simp add: card_word of_nat_inj)
  moreover have "{0..<n}  {i. i < CARD('a word)}"
    using that by auto
  ultimately have "inj_on ?of_nat {0..<n}"
    by (rule inj_on_subset)
  then show ?thesis
    by (simp add: card_image)
qed

lemma card_map_elide2:
  "n  CARD('a::len word)  card ((of_nat::nat  'a::len word) ` {0..<n}) = n"
  by (subst card_map_elide) clarsimp+

lemma eq_ucast_ucast_eq:
  "LENGTH('b)  LENGTH('a)  x = ucast y  ucast x = y"
  for x :: "'a::len word" and y :: "'b::len word"
  by transfer simp

lemma le_ucast_ucast_le:
  "x  ucast y  ucast x  y"
  for x :: "'a::len word" and y :: "'b::len word"
  by (smt (verit) le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1))

lemma less_ucast_ucast_less:
  "LENGTH('b)  LENGTH('a)  x < ucast y  ucast x < y"
  for x :: "'a::len word" and y :: "'b::len word"
  by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less)

lemma ucast_le_ucast:
  "LENGTH('a)  LENGTH('b)  (ucast x  (ucast y::'b::len word)) = (x  y)"
  for x :: "'a::len word"
  by (simp add: unat_arith_simps(1) unat_ucast_up_simp)

lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2]

lemma ucast_or_distrib:
  fixes x :: "'a::len word"
  fixes y :: "'a::len word"
  shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y"
  by (fact unsigned_or_eq)

lemma word_exists_nth:
  "(w::'a::len word)  0  i. bit w i"
  by (auto simp: bit_eq_iff)

lemma max_word_not_0 [simp]:
  "- 1  (0 :: 'a::len word)"
  by simp

lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)"
  using unat_gt_0 [of - 1 :: 'a::len word] by simp


(* Miscellaneous conditional injectivity rules. *)

lemma mult_pow2_inj:
  assumes ws: "m + n  LENGTH('a)"
  assumes le: "x  mask m" "y  mask m"
  assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)"
  shows "x = y"
proof (rule bit_word_eqI)
  fix q
  assume q < LENGTH('a)
  from eq have push_bit n x = push_bit n y
    by (simp add: push_bit_eq_mult)
  moreover from le have §: take_bit m x = x take_bit m y = y
    by (simp_all add: less_eq_mask_iff_take_bit_eq_self)
  ultimately have push_bit n (take_bit m x) = push_bit n (take_bit m y)
    by simp_all
  with § q < LENGTH('a) ws show bit x q  bit y q
    apply (simp add: push_bit_take_bit bit_eq_iff bit_simps not_le)
    by (metis (full_types) add.commute add_diff_cancel_right' add_less_cancel_right le_add2 less_le_trans)
qed

lemma word_of_nat_inj:
  assumes "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)"
  assumes "of_nat x = (of_nat y :: 'a::len word)"
  shows "x = y"
  using assms of_nat_inj by blast

lemma word_of_int_bin_cat_eq_iff:
  "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) =
    word_of_int (concat_bit LENGTH('b) (uint d) (uint c))  b = d  a = c" (is "?L=?R")
  if "LENGTH('a) + LENGTH('b)  LENGTH('c)"
  for a:: "'a::len word" and b:: "'b::len word"
proof
  assume ?L
  then have "take_bit LENGTH('c) (concat_bit LENGTH('b) (uint b) (uint a)) =
             take_bit LENGTH('c) (concat_bit LENGTH('b) (uint d) (uint c))"
    by (simp add: word_of_int_eq_iff)
  then show "b = d  a = c"
    using that concat_bit_eq_iff 
    by (metis (no_types, lifting) concat_bit_assoc concat_bit_of_zero_2 concat_bit_take_bit_eq 
        take_bit_tightened uint_sint unsigned_word_eqI)
qed auto

lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d  a = c  b = d"
  if "LENGTH('a) + LENGTH('b)  LENGTH('c)"
  for a:: "'a::len word" and b:: "'b::len word"
  using word_of_int_bin_cat_eq_iff [OF that, of b a d c]
  by (simp add: word_cat_eq' ac_simps)

lemma p2_eq_1: "2 ^ n = (1::'a::len word)  n = 0"
proof -
  have "2 ^ n = (1::'a word)  n = 0"
    by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0
        power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one)
  then show ?thesis by auto
qed

end

lemmas word_div_less = div_word_less

lemma word_mod_by_0: "k mod (0::'a::len word) = k"
  by (simp add: word_arith_nat_mod)

― ‹the binary operations only›  (* BH: why is this needed? *)
lemmas word_log_binary_defs =
  word_and_def word_or_def word_xor_def

― ‹limited hom result›
lemma word_cat_hom:
  "LENGTH('a::len)  LENGTH('b::len) + LENGTH('c::len) 
    (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) =
    word_of_int ((λk n l. concat_bit n l k) w (size b) (uint b))"
  by (rule bit_eqI) (auto simp: bit_simps size_word_def)

lemma uint_shiftl:
  "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))"
  by (simp add: unsigned_push_bit_eq word_size)

lemma sint_range':
  - (2 ^ (LENGTH('a) - Suc 0))  sint x  sint x < 2 ^ (LENGTH('a) - Suc 0)
  for x :: 'a::len word
  by transfer simp_all

lemma signed_arith_eq_checks_to_ord:
  sint a + sint b = sint (a + b)  (a ≤s a + b  0 ≤s b) (is ?P)
  sint a - sint b = sint (a - b)  (0 ≤s a - b  b ≤s a) (is ?Q)
  - sint a = sint (- a)  (0 ≤s - a  a ≤s 0) (is ?R)
proof -
  define n where n = LENGTH('a) - Suc 0
  then have [simp]: LENGTH('a) = Suc n
    by simp
  define k l where k = sint a l = sint b
  then have [simp]: sint a = k sint b = l
    by simp_all
  have self_eq_iff: k + l = signed_take_bit n (k + l)  - (2 ^ n)  k + l  k + l < 2 ^ n
    using signed_take_bit_int_eq_self_iff [of n k + l]
    by auto
  from sint_range' [where x=a] sint_range' [where x=b]
  have assms: - (2 ^ n)  k k < 2 ^ n - (2 ^ n)  l l < 2 ^ n
    by simp_all
  have aux: signed_take_bit n x 
           = (if x < - (2 ^ n) then x + 2 * (2 ^ n)
              else if x  2 ^ n then x - 2 * (2 ^ n) else x)
    if x: - 3 * (2 ^ n)  x  x < 3 * (2 ^ n) for x :: int
  proof -
    have "2 ^ n + (x + 2 ^ n) mod (2 * 2 ^ n) = x"
      if "x < 3 * 2 ^ n" "2 ^ n  x"
      using that by (smt (verit) minus_mod_self2 mod_pos_pos_trivial)
    moreover have "(x + 2 ^ n) mod (2 * 2 ^ n) = 3 * 2 ^ n + x"
      if "- (3 * 2 ^ n)  x" and "x < - (2 ^ n)"
      using that by (smt (verit) mod_add_self2 mod_pos_pos_trivial)
    ultimately show ?thesis
      using x by (auto simp: signed_take_bit_eq_take_bit_shift take_bit_eq_mod)
  qed
  show ?P ?Q ?R
    using assms
    by (auto simp: sint_word_ariths word_sle_eq word_sless_alt aux)
qed

lemma signed_mult_eq_checks_double_size:
  assumes mult_le: "(2 ^ (LENGTH('a) - 1) + 1) ^ 2  (2 :: int) ^ (LENGTH('b) - 1)"
           and le: "2 ^ (LENGTH('a) - 1)  (2 :: int) ^ (LENGTH('b) - 1)"
  shows "(sint (a :: 'a :: len word) * sint b = sint (a * b))
       = (scast a * scast b = (scast (a * b) :: 'b :: len word))"
proof -
  have P: "signed_take_bit (size a - 1) (sint a * sint b)  range (signed_take_bit (size a - 1))"
    by simp

  have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1"
    by (smt (verit) sint_ge sint_lt wsst_TYs(3))
  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)
  with abs_ab le show ?thesis
    by (smt (verit) One_nat_def Word.of_int_sint of_int_mult sint_greater_eq sint_less sint_of_int_eq)
qed

context
  includes bit_operations_syntax
begin

lemma wils1:
  word_of_int (NOT (uint (word_of_int x :: 'a word))) = (word_of_int (NOT x) :: 'a::len word)
  word_of_int (uint (word_of_int x :: 'a word) XOR uint (word_of_int y :: 'a word)) = (word_of_int (x XOR y) :: 'a::len word)
  word_of_int (uint (word_of_int x :: 'a word) AND uint (word_of_int y :: 'a word)) = (word_of_int (x AND y) :: 'a::len word)
  word_of_int (uint (word_of_int x :: 'a word) OR uint (word_of_int y :: 'a word)) = (word_of_int (x OR y) :: 'a::len word)
  by (simp_all add: word_of_int_eq_iff uint_word_of_int_eq take_bit_not_take_bit)

end

end