src/HOL/Word/WordBitwise.thy
author nipkow
Wed, 10 Jan 2018 15:25:09 +0100
changeset 67399 eab6ce8368fa
parent 67121 116968454d70
child 67405 e9ab4ad7bd15
permissions -rw-r--r--
ran isabelle update_op on all sources
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     1
(*  Title:      HOL/Word/WordBitwise.thy
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     2
    Authors:    Thomas Sewell, NICTA and Sascha Boehme, TU Muenchen
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     3
*)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     4
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     5
theory WordBitwise
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
     6
  imports Word
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     7
begin
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
     8
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61144
diff changeset
     9
text \<open>Helper constants used in defining addition\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    10
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    11
definition xor3 :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    12
  where "xor3 a b c = (a = (b = c))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    13
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    14
definition carry :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    15
  where "carry a b c = ((a \<and> (b \<or> c)) \<or> (b \<and> c))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    16
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    17
lemma carry_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    18
  "carry True a b = (a \<or> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    19
  "carry a True b = (a \<or> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    20
  "carry a b True = (a \<or> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    21
  "carry False a b = (a \<and> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    22
  "carry a False b = (a \<and> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    23
  "carry a b False = (a \<and> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    24
  by (auto simp add: carry_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    25
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    26
lemma xor3_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    27
  "xor3 True a b = (a = b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    28
  "xor3 a True b = (a = b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    29
  "xor3 a b True = (a = b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    30
  "xor3 False a b = (a \<noteq> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    31
  "xor3 a False b = (a \<noteq> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    32
  "xor3 a b False = (a \<noteq> b)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    33
  by (simp_all add: xor3_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    34
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61144
diff changeset
    35
text \<open>Breaking up word equalities into equalities on their
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    36
  bit lists. Equalities are generated and manipulated in the
66446
aeb8b8fe94d0 fix document
Lars Hupel <lars.hupel@mytum.de>
parents: 65363
diff changeset
    37
  reverse order to @{const to_bl}.\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    38
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    39
lemma word_eq_rbl_eq: "x = y \<longleftrightarrow> rev (to_bl x) = rev (to_bl y)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    40
  by simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    41
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
    42
lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\<or>) (rev (to_bl x)) (rev (to_bl y))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    43
  by (simp add: map2_def zip_rev bl_word_or rev_map)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    44
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
    45
lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\<and>) (rev (to_bl x)) (rev (to_bl y))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    46
  by (simp add: map2_def zip_rev bl_word_and rev_map)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    47
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
    48
lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\<noteq>) (rev (to_bl x)) (rev (to_bl y))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    49
  by (simp add: map2_def zip_rev bl_word_xor rev_map)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    50
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    51
lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    52
  by (simp add: bl_word_not rev_map)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    53
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    54
lemma bl_word_sub: "to_bl (x - y) = to_bl (x + (- y))"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 51717
diff changeset
    55
  by simp
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    56
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    57
lemma rbl_word_1: "rev (to_bl (1 :: 'a::len0 word)) = takefill False (len_of TYPE('a)) [True]"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    58
  apply (rule_tac s="rev (to_bl (word_succ (0 :: 'a word)))" in trans)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    59
   apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    60
  apply (simp only: rtb_rbl_ariths(1)[OF refl])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    61
  apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    62
  apply (case_tac "len_of TYPE('a)")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    63
   apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    64
  apply (simp add: takefill_alt)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    65
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    66
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    67
lemma rbl_word_if: "rev (to_bl (if P then x else y)) = map2 (If P) (rev (to_bl x)) (rev (to_bl y))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    68
  by (simp add: map2_def split_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    69
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    70
lemma rbl_add_carry_Cons:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    71
  "(if car then rbl_succ else id) (rbl_add (x # xs) (y # ys)) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    72
    xor3 x y car # (if carry x y car then rbl_succ else id) (rbl_add xs ys)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    73
  by (simp add: carry_def xor3_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    74
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    75
lemma rbl_add_suc_carry_fold:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    76
  "length xs = length ys \<Longrightarrow>
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    77
    \<forall>car. (if car then rbl_succ else id) (rbl_add xs ys) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    78
      (foldr (\<lambda>(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\<lambda>_. [])) car"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    79
  apply (erule list_induct2)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    80
   apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    81
  apply (simp only: rbl_add_carry_Cons)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    82
  apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    83
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    84
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    85
lemma to_bl_plus_carry:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    86
  "to_bl (x + y) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    87
    rev (foldr (\<lambda>(x, y) res car. xor3 x y car # res (carry x y car))
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    88
      (rev (zip (to_bl x) (to_bl y))) (\<lambda>_. []) False)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    89
  using rbl_add_suc_carry_fold[where xs="rev (to_bl x)" and ys="rev (to_bl y)"]
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    90
  apply (simp add: word_add_rbl[OF refl refl])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    91
  apply (drule_tac x=False in spec)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    92
  apply (simp add: zip_rev)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    93
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    94
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    95
definition "rbl_plus cin xs ys =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    96
  foldr (\<lambda>(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\<lambda>_. []) cin"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    97
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
    98
lemma rbl_plus_simps:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
    99
  "rbl_plus cin (x # xs) (y # ys) = xor3 x y cin # rbl_plus (carry x y cin) xs ys"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   100
  "rbl_plus cin [] ys = []"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   101
  "rbl_plus cin xs [] = []"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   102
  by (simp_all add: rbl_plus_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   103
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   104
lemma rbl_word_plus: "rev (to_bl (x + y)) = rbl_plus False (rev (to_bl x)) (rev (to_bl y))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   105
  by (simp add: rbl_plus_def to_bl_plus_carry zip_rev)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   106
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   107
definition "rbl_succ2 b xs = (if b then rbl_succ xs else xs)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   108
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   109
lemma rbl_succ2_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   110
  "rbl_succ2 b [] = []"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   111
  "rbl_succ2 b (x # xs) = (b \<noteq> x) # rbl_succ2 (x \<and> b) xs"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   112
  by (simp_all add: rbl_succ2_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   113
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   114
lemma twos_complement: "- x = word_succ (NOT x)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   115
  using arg_cong[OF word_add_not[where x=x], where f="\<lambda>a. a - x + 1"]
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   116
  by (simp add: word_succ_p1 word_sp_01[unfolded word_succ_p1] del: word_add_not)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   117
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   118
lemma rbl_word_neg: "rev (to_bl (- x)) = rbl_succ2 True (map Not (rev (to_bl x)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   119
  by (simp add: twos_complement word_succ_rbl[OF refl] bl_word_not rev_map rbl_succ2_def)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   120
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   121
lemma rbl_word_cat:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   122
  "rev (to_bl (word_cat x y :: 'a::len0 word)) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   123
    takefill False (len_of TYPE('a)) (rev (to_bl y) @ rev (to_bl x))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   124
  by (simp add: word_cat_bl word_rev_tf)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   125
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   126
lemma rbl_word_slice:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   127
  "rev (to_bl (slice n w :: 'a::len0 word)) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   128
    takefill False (len_of TYPE('a)) (drop n (rev (to_bl w)))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   129
  apply (simp add: slice_take word_rev_tf rev_take)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   130
  apply (cases "n < len_of TYPE('b)", simp_all)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   131
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   132
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   133
lemma rbl_word_ucast:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   134
  "rev (to_bl (ucast x :: 'a::len0 word)) = takefill False (len_of TYPE('a)) (rev (to_bl x))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   135
  apply (simp add: to_bl_ucast takefill_alt)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   136
  apply (simp add: rev_drop)
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   137
  apply (cases "len_of TYPE('a) < len_of TYPE('b)")
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   138
   apply simp_all
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   139
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   140
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   141
lemma rbl_shiftl:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   142
  "rev (to_bl (w << n)) = takefill False (size w) (replicate n False @ rev (to_bl w))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   143
  by (simp add: bl_shiftl takefill_alt word_size rev_drop)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   144
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   145
lemma rbl_shiftr:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   146
  "rev (to_bl (w >> n)) = takefill False (size w) (drop n (rev (to_bl w)))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   147
  by (simp add: shiftr_slice rbl_word_slice word_size)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   148
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   149
definition "drop_nonempty v n xs = (if n < length xs then drop n xs else [last (v # xs)])"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   150
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   151
lemma drop_nonempty_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   152
  "drop_nonempty v (Suc n) (x # xs) = drop_nonempty x n xs"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   153
  "drop_nonempty v 0 (x # xs) = (x # xs)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   154
  "drop_nonempty v n [] = [v]"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   155
  by (simp_all add: drop_nonempty_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   156
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   157
definition "takefill_last x n xs = takefill (last (x # xs)) n xs"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   158
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   159
lemma takefill_last_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   160
  "takefill_last z (Suc n) (x # xs) = x # takefill_last x n xs"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   161
  "takefill_last z 0 xs = []"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   162
  "takefill_last z n [] = replicate n z"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   163
  by (simp_all add: takefill_last_def) (simp_all add: takefill_alt)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   164
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   165
lemma rbl_sshiftr:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   166
  "rev (to_bl (w >>> n)) = takefill_last False (size w) (drop_nonempty False n (rev (to_bl w)))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   167
  apply (cases "n < size w")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   168
   apply (simp add: bl_sshiftr takefill_last_def word_size
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   169
      takefill_alt rev_take last_rev
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   170
      drop_nonempty_def)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   171
  apply (subgoal_tac "(w >>> n) = of_bl (replicate (size w) (msb w))")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   172
   apply (simp add: word_size takefill_last_def takefill_alt
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   173
      last_rev word_msb_alt word_rev_tf
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   174
      drop_nonempty_def take_Cons')
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   175
   apply (case_tac "len_of TYPE('a)", simp_all)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   176
  apply (rule word_eqI)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   177
  apply (simp add: nth_sshiftr word_size test_bit_of_bl
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   178
      msb_nth)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   179
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   180
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   181
lemma nth_word_of_int:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   182
  "(word_of_int x :: 'a::len0 word) !! n = (n < len_of TYPE('a) \<and> bin_nth x n)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   183
  apply (simp add: test_bit_bl word_size to_bl_of_bin)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   184
  apply (subst conj_cong[OF refl], erule bin_nth_bl)
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   185
  apply auto
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   186
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   187
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   188
lemma nth_scast:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   189
  "(scast (x :: 'a::len word) :: 'b::len word) !! n =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   190
    (n < len_of TYPE('b) \<and>
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   191
    (if n < len_of TYPE('a) - 1 then x !! n
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   192
     else x !! (len_of TYPE('a) - 1)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   193
  by (simp add: scast_def nth_sint)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   194
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   195
lemma rbl_word_scast:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   196
  "rev (to_bl (scast x :: 'a::len word)) = takefill_last False (len_of TYPE('a)) (rev (to_bl x))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   197
  apply (rule nth_equalityI)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   198
   apply (simp add: word_size takefill_last_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   199
  apply (clarsimp simp: nth_scast takefill_last_def
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   200
      nth_takefill word_size nth_rev to_bl_nth)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   201
  apply (cases "len_of TYPE('b)")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   202
   apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   203
  apply (clarsimp simp: less_Suc_eq_le linorder_not_less
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   204
      last_rev word_msb_alt[symmetric]
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   205
      msb_nth)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   206
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   207
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   208
definition rbl_mul :: "bool list \<Rightarrow> bool list \<Rightarrow> bool list"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   209
  where "rbl_mul xs ys = foldr (\<lambda>x sm. rbl_plus False (map ((\<and>) x) ys) (False # sm)) xs []"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   210
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   211
lemma rbl_mul_simps:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   212
  "rbl_mul (x # xs) ys = rbl_plus False (map ((\<and>) x) ys) (False # rbl_mul xs ys)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   213
  "rbl_mul [] ys = []"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   214
  by (simp_all add: rbl_mul_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   215
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   216
lemma takefill_le2: "length xs \<le> n \<Longrightarrow> takefill x m (takefill x n xs) = takefill x m xs"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   217
  by (simp add: takefill_alt replicate_add[symmetric])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   218
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   219
lemma take_rbl_plus: "\<forall>n b. take n (rbl_plus b xs ys) = rbl_plus b (take n xs) (take n ys)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   220
  apply (simp add: rbl_plus_def take_zip[symmetric])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   221
  apply (rule_tac list="zip xs ys" in list.induct)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   222
   apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   223
  apply (clarsimp simp: split_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   224
  apply (case_tac n, simp_all)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   225
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   226
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   227
lemma word_rbl_mul_induct:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   228
  "length xs \<le> size y \<Longrightarrow>
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   229
    rbl_mul xs (rev (to_bl y)) = take (length xs) (rev (to_bl (of_bl (rev xs) * y)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   230
  for y :: "'a::len word"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   231
proof (induct xs)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   232
  case Nil
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   233
  show ?case by (simp add: rbl_mul_simps)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   234
next
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   235
  case (Cons z zs)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   236
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   237
  have rbl_word_plus': "to_bl (x + y) = rev (rbl_plus False (rev (to_bl x)) (rev (to_bl y)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   238
    for x y :: "'a word"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   239
    by (simp add: rbl_word_plus[symmetric])
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   240
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   241
  have mult_bit: "to_bl (of_bl [z] * y) = map ((\<and>) z) (to_bl y)"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   242
    by (cases z) (simp cong: map_cong, simp add: map_replicate_const cong: map_cong)
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   243
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   244
  have shiftl: "of_bl xs * 2 * y = (of_bl xs * y) << 1" for xs
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   245
    by (simp add: shiftl_t2n)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   246
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   247
  have zip_take_triv: "\<And>xs ys n. n = length ys \<Longrightarrow> zip (take n xs) ys = zip xs ys"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   248
    by (rule nth_equalityI) simp_all
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   249
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   250
  from Cons show ?case
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55465
diff changeset
   251
    apply (simp add: trans [OF of_bl_append add.commute]
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   252
        rbl_mul_simps rbl_word_plus' distrib_right mult_bit shiftl rbl_shiftl)
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   253
    apply (simp add: takefill_alt word_size rev_map take_rbl_plus min_def)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   254
    apply (simp add: rbl_plus_def zip_take_triv)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   255
    done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   256
qed
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   257
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   258
lemma rbl_word_mul: "rev (to_bl (x * y)) = rbl_mul (rev (to_bl x)) (rev (to_bl y))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   259
  for x :: "'a::len word"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   260
  using word_rbl_mul_induct[where xs="rev (to_bl x)" and y=y] by (simp add: word_size)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   261
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61144
diff changeset
   262
text \<open>Breaking up inequalities into bitlist properties.\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   263
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   264
definition
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   265
  "rev_bl_order F xs ys =
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   266
     (length xs = length ys \<and>
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   267
       ((xs = ys \<and> F)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   268
          \<or> (\<exists>n < length xs. drop (Suc n) xs = drop (Suc n) ys
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   269
                   \<and> \<not> xs ! n \<and> ys ! n)))"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   270
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   271
lemma rev_bl_order_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   272
  "rev_bl_order F [] [] = F"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   273
  "rev_bl_order F (x # xs) (y # ys) = rev_bl_order ((y \<and> \<not> x) \<or> ((y \<or> \<not> x) \<and> F)) xs ys"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   274
   apply (simp_all add: rev_bl_order_def)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   275
  apply (rule conj_cong[OF refl])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   276
  apply (cases "xs = ys")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   277
   apply (simp add: nth_Cons')
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   278
   apply blast
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   279
  apply (simp add: nth_Cons')
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   280
  apply safe
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   281
   apply (rule_tac x="n - 1" in exI)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   282
   apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   283
  apply (rule_tac x="Suc n" in exI)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   284
  apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   285
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   286
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   287
lemma rev_bl_order_rev_simp:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   288
  "length xs = length ys \<Longrightarrow>
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   289
   rev_bl_order F (xs @ [x]) (ys @ [y]) = ((y \<and> \<not> x) \<or> ((y \<or> \<not> x) \<and> rev_bl_order F xs ys))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   290
  by (induct arbitrary: F rule: list_induct2) (auto simp: rev_bl_order_simps)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   291
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   292
lemma rev_bl_order_bl_to_bin:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   293
  "length xs = length ys \<Longrightarrow>
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   294
    rev_bl_order True xs ys = (bl_to_bin (rev xs) \<le> bl_to_bin (rev ys)) \<and>
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   295
    rev_bl_order False xs ys = (bl_to_bin (rev xs) < bl_to_bin (rev ys))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   296
  apply (induct xs ys rule: list_induct2)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   297
   apply (simp_all add: rev_bl_order_simps bl_to_bin_app_cat)
54847
d6cf9a5b9be9 prefer plain bool over dedicated type for binary digits
haftmann
parents: 54489
diff changeset
   298
  apply (auto simp add: bl_to_bin_def Bit_B0 Bit_B1 add1_zle_eq Bit_def)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   299
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   300
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   301
lemma word_le_rbl: "x \<le> y \<longleftrightarrow> rev_bl_order True (rev (to_bl x)) (rev (to_bl y))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   302
  for x y :: "'a::len0 word"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   303
  by (simp add: rev_bl_order_bl_to_bin word_le_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   304
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   305
lemma word_less_rbl: "x < y \<longleftrightarrow> rev_bl_order False (rev (to_bl x)) (rev (to_bl y))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   306
  for x y :: "'a::len0 word"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   307
  by (simp add: word_less_alt rev_bl_order_bl_to_bin)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   308
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   309
lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   310
  apply (cases "msb x")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   311
   apply (rule word_sint.Abs_eqD[where 'a='a], simp_all)
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   312
    apply (simp add: word_size wi_hom_syms word_of_int_2p_len)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   313
   apply (simp add: sints_num word_size)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   314
   apply (rule conjI)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   315
    apply (simp add: le_diff_eq')
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   316
    apply (rule order_trans[where y="2 ^ (len_of TYPE('a) - 1)"])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   317
     apply (simp add: power_Suc[symmetric])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   318
    apply (simp add: linorder_not_less[symmetric] mask_eq_iff[symmetric])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   319
    apply (rule notI, drule word_eqD[where x="size x - 1"])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   320
    apply (simp add: msb_nth word_ops_nth_size word_size)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   321
   apply (simp add: order_less_le_trans[where y=0])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   322
  apply (rule word_uint.Abs_eqD[where 'a='a], simp_all)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   323
  apply (simp add: linorder_not_less uints_num word_msb_sint)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   324
  apply (rule order_less_le_trans[OF sint_lt])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   325
  apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   326
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   327
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   328
lemma word_sle_msb_le: "x <=s y \<longleftrightarrow> (msb y \<longrightarrow> msb x) \<and> ((msb x \<and> \<not> msb y) \<or> x \<le> y)"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   329
  apply (simp add: word_sle_def word_sint_msb_eq word_size word_le_def)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   330
  apply safe
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   331
   apply (rule order_trans[OF _ uint_ge_0])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   332
   apply (simp add: order_less_imp_le)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   333
  apply (erule notE[OF leD])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   334
  apply (rule order_less_le_trans[OF _ uint_ge_0])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   335
  apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   336
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   337
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   338
lemma word_sless_msb_less: "x <s y \<longleftrightarrow> (msb y \<longrightarrow> msb x) \<and> ((msb x \<and> \<not> msb y) \<or> x < y)"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   339
  by (auto simp add: word_sless_def word_sle_msb_le)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   340
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   341
definition "map_last f xs = (if xs = [] then [] else butlast xs @ [f (last xs)])"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   342
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   343
lemma map_last_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   344
  "map_last f [] = []"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   345
  "map_last f [x] = [f x]"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   346
  "map_last f (x # y # zs) = x # map_last f (y # zs)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   347
  by (simp_all add: map_last_def)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   348
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   349
lemma word_sle_rbl:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   350
  "x <=s y \<longleftrightarrow> rev_bl_order True (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   351
  using word_msb_alt[where w=x] word_msb_alt[where w=y]
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   352
  apply (simp add: word_sle_msb_le word_le_rbl)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   353
  apply (subgoal_tac "length (to_bl x) = length (to_bl y)")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   354
   apply (cases "to_bl x", simp)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   355
   apply (cases "to_bl y", simp)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   356
   apply (clarsimp simp: map_last_def rev_bl_order_rev_simp)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   357
   apply auto
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   358
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   359
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   360
lemma word_sless_rbl:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   361
  "x <s y \<longleftrightarrow> rev_bl_order False (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   362
  using word_msb_alt[where w=x] word_msb_alt[where w=y]
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   363
  apply (simp add: word_sless_msb_less word_less_rbl)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   364
  apply (subgoal_tac "length (to_bl x) = length (to_bl y)")
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   365
   apply (cases "to_bl x", simp)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   366
   apply (cases "to_bl y", simp)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   367
   apply (clarsimp simp: map_last_def rev_bl_order_rev_simp)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   368
   apply auto
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   369
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   370
66446
aeb8b8fe94d0 fix document
Lars Hupel <lars.hupel@mytum.de>
parents: 65363
diff changeset
   371
text \<open>Lemmas for unpacking @{term "rev (to_bl n)"} for numerals n and also
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   372
  for irreducible values and expressions.\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   373
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   374
lemma rev_bin_to_bl_simps:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   375
  "rev (bin_to_bl 0 x) = []"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   376
  "rev (bin_to_bl (Suc n) (numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (numeral nm))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   377
  "rev (bin_to_bl (Suc n) (numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (numeral nm))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   378
  "rev (bin_to_bl (Suc n) (numeral (num.One))) = True # replicate n False"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   379
  "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (- numeral nm))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   380
  "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm))) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   381
    True # rev (bin_to_bl n (- numeral (nm + num.One)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   382
  "rev (bin_to_bl (Suc n) (- numeral (num.One))) = True # replicate n True"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   383
  "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm + num.One))) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   384
    True # rev (bin_to_bl n (- numeral (nm + num.One)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   385
  "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm + num.One))) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   386
    False # rev (bin_to_bl n (- numeral (nm + num.One)))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   387
  "rev (bin_to_bl (Suc n) (- numeral (num.One + num.One))) =
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   388
    False # rev (bin_to_bl n (- numeral num.One))"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   389
  apply simp_all
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   390
  apply (simp_all only: bin_to_bl_aux_alt)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   391
  apply (simp_all)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   392
  apply (simp_all add: bin_to_bl_zero_aux bin_to_bl_minus1_aux)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   393
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   394
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   395
lemma to_bl_upt: "to_bl x = rev (map ((!!) x) [0 ..< size x])"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   396
  apply (rule nth_equalityI)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   397
   apply (simp add: word_size)
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   398
  apply (auto simp: to_bl_nth word_size nth_rev)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   399
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   400
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   401
lemma rev_to_bl_upt: "rev (to_bl x) = map ((!!) x) [0 ..< size x]"
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   402
  by (simp add: to_bl_upt)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   403
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   404
lemma upt_eq_list_intros:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   405
  "j \<le> i \<Longrightarrow> [i ..< j] = []"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   406
  "i = x \<Longrightarrow> x < j \<Longrightarrow> [x + 1 ..< j] = xs \<Longrightarrow> [i ..< j] = (x # xs)"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   407
  by (simp_all add: upt_eq_Cons_conv)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   408
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   409
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   410
subsection \<open>Tactic definition\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   411
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61144
diff changeset
   412
ML \<open>
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   413
structure Word_Bitwise_Tac =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   414
struct
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   415
51686
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 50902
diff changeset
   416
val word_ss = simpset_of @{theory_context Word};
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   417
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   418
fun mk_nat_clist ns =
67121
wenzelm
parents: 67120
diff changeset
   419
  fold_rev (Thm.mk_binop @{cterm "Cons :: nat \<Rightarrow> _"})
wenzelm
parents: 67120
diff changeset
   420
    ns @{cterm "[] :: nat list"};
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   421
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   422
fun upt_conv ctxt ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   423
  case Thm.term_of ct of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   424
    (@{const upt} $ n $ m) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   425
      let
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 57512
diff changeset
   426
        val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   427
        val ns = map (Numeral.mk_cnumber @{ctyp nat}) (i upto (j - 1))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   428
          |> mk_nat_clist;
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   429
        val prop =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   430
          Thm.mk_binop @{cterm "(=) :: nat list \<Rightarrow> _"} ct ns
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   431
          |> Thm.apply @{cterm Trueprop};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   432
      in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   433
        try (fn () =>
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 62913
diff changeset
   434
          Goal.prove_internal ctxt [] prop
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   435
            (K (REPEAT_DETERM (resolve_tac ctxt @{thms upt_eq_list_intros} 1
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   436
                ORELSE simp_tac (put_simpset word_ss ctxt) 1))) |> mk_meta_eq) ()
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   437
      end
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   438
  | _ => NONE;
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   439
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   440
val expand_upt_simproc =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   441
  Simplifier.make_simproc @{context} "expand_upt"
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62390
diff changeset
   442
   {lhss = [@{term "upt x y"}], proc = K upt_conv};
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   443
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   444
fun word_len_simproc_fn ctxt ct =
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   445
  (case Thm.term_of ct of
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   446
    Const (@{const_name len_of}, _) $ t =>
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   447
     (let
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   448
        val T = fastype_of t |> dest_Type |> snd |> the_single
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   449
        val n = Numeral.mk_cnumber @{ctyp nat} (Word_Lib.dest_binT T);
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   450
        val prop =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   451
          Thm.mk_binop @{cterm "(=) :: nat \<Rightarrow> _"} ct n
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   452
          |> Thm.apply @{cterm Trueprop};
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   453
      in
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   454
        Goal.prove_internal ctxt [] prop (K (simp_tac (put_simpset word_ss ctxt) 1))
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   455
        |> mk_meta_eq |> SOME
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   456
      end handle TERM _ => NONE | TYPE _ => NONE)
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   457
  | _ => NONE);
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   458
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   459
val word_len_simproc =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   460
  Simplifier.make_simproc @{context} "word_len"
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62390
diff changeset
   461
   {lhss = [@{term "len_of x"}], proc = K word_len_simproc_fn};
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   462
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   463
(* convert 5 or nat 5 to Suc 4 when n_sucs = 1, Suc (Suc 4) when n_sucs = 2,
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   464
   or just 5 (discarding nat) when n_sucs = 0 *)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   465
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   466
fun nat_get_Suc_simproc_fn n_sucs ctxt ct =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   467
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   468
    val (f $ arg) = Thm.term_of ct;
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   469
    val n =
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   470
      (case arg of @{term nat} $ n => n | n => n)
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   471
      |> HOLogic.dest_number |> snd;
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   472
    val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs) else (n, 0);
67121
wenzelm
parents: 67120
diff changeset
   473
    val arg' = funpow i HOLogic.mk_Suc (HOLogic.mk_number @{typ nat} j);
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   474
    val _ = if arg = arg' then raise TERM ("", []) else ();
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   475
    fun propfn g =
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   476
      HOLogic.mk_eq (g arg, g arg')
59643
f3be9235503d clarified context;
wenzelm
parents: 59621
diff changeset
   477
      |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt;
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   478
    val eq1 =
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   479
      Goal.prove_internal ctxt [] (propfn I)
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   480
        (K (simp_tac (put_simpset word_ss ctxt) 1));
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   481
  in
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   482
    Goal.prove_internal ctxt [] (propfn (curry ($) f))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   483
      (K (simp_tac (put_simpset HOL_ss ctxt addsimps [eq1]) 1))
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   484
    |> mk_meta_eq |> SOME
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66446
diff changeset
   485
  end handle TERM _ => NONE;
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   486
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   487
fun nat_get_Suc_simproc n_sucs ts =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   488
  Simplifier.make_simproc @{context} "nat_get_Suc"
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   489
   {lhss = map (fn t => t $ @{term "n :: nat"}) ts,
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62390
diff changeset
   490
    proc = K (nat_get_Suc_simproc_fn n_sucs)};
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   491
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   492
val no_split_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   493
  simpset_of (put_simpset HOL_ss @{context}
62390
842917225d56 more canonical names
nipkow
parents: 61799
diff changeset
   494
    |> Splitter.del_split @{thm if_split});
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   495
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   496
val expand_word_eq_sss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   497
  (simpset_of (put_simpset HOL_basic_ss @{context} addsimps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   498
       @{thms word_eq_rbl_eq word_le_rbl word_less_rbl word_sle_rbl word_sless_rbl}),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   499
  map simpset_of [
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   500
   put_simpset no_split_ss @{context} addsimps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   501
    @{thms rbl_word_plus rbl_word_and rbl_word_or rbl_word_not
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   502
                                rbl_word_neg bl_word_sub rbl_word_xor
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   503
                                rbl_word_cat rbl_word_slice rbl_word_scast
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   504
                                rbl_word_ucast rbl_shiftl rbl_shiftr rbl_sshiftr
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   505
                                rbl_word_if},
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   506
   put_simpset no_split_ss @{context} addsimps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   507
    @{thms to_bl_numeral to_bl_neg_numeral to_bl_0 rbl_word_1},
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   508
   put_simpset no_split_ss @{context} addsimps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   509
    @{thms rev_rev_ident rev_replicate rev_map to_bl_upt word_size}
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   510
          addsimprocs [word_len_simproc],
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   511
   put_simpset no_split_ss @{context} addsimps
55465
0d31c0546286 merged 'List.map' and 'List.list.map'
blanchet
parents: 54883
diff changeset
   512
    @{thms list.simps split_conv replicate.simps list.map
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   513
                                zip_Cons_Cons zip_Nil drop_Suc_Cons drop_0 drop_Nil
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   514
                                foldr.simps map2_Cons map2_Nil takefill_Suc_Cons
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   515
                                takefill_Suc_Nil takefill.Z rbl_succ2_simps
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   516
                                rbl_plus_simps rev_bin_to_bl_simps append.simps
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   517
                                takefill_last_simps drop_nonempty_simps
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   518
                                rev_bl_order_simps}
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   519
          addsimprocs [expand_upt_simproc,
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54847
diff changeset
   520
                       nat_get_Suc_simproc 4
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   521
                         [@{term replicate}, @{term "takefill x"},
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   522
                          @{term drop}, @{term "bin_to_bl"},
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   523
                          @{term "takefill_last x"},
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 59643
diff changeset
   524
                          @{term "drop_nonempty x"}]],
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   525
    put_simpset no_split_ss @{context} addsimps @{thms xor3_simps carry_simps if_bool_simps}
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   526
  ])
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   527
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   528
fun tac ctxt =
50107
289181e3e524 tuned signature;
wenzelm
parents: 49962
diff changeset
   529
  let
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   530
    val (ss, sss) = expand_word_eq_sss;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   531
  in
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67121
diff changeset
   532
    foldr1 (THEN_ALL_NEW)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   533
      ((CHANGED o safe_full_simp_tac (put_simpset ss ctxt)) ::
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   534
        map (fn ss => safe_full_simp_tac (put_simpset ss ctxt)) sss)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   535
  end;
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   536
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   537
end
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61144
diff changeset
   538
\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   539
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   540
method_setup word_bitwise =
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61144
diff changeset
   541
  \<open>Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (Word_Bitwise_Tac.tac ctxt 1))\<close>
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   542
  "decomposer for word equalities and inequalities into bit propositions"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   543
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
diff changeset
   544
end