src/HOL/Library/Word.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 82687 97b9ac57751e
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
72515
c7038c397ae3 moved most material from session HOL-Word to Word_Lib in the AFP
haftmann
parents: 72512
diff changeset
     1
(*  Title:      HOL/Library/Word.thy
c7038c397ae3 moved most material from session HOL-Word to Word_Lib in the AFP
haftmann
parents: 72512
diff changeset
     2
    Author:     Jeremy Dawson and Gerwin Klein, NICTA, et. al.
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     3
*)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     4
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
     5
section \<open>A type of finite bit strings\<close>
24350
4d74f37c6367 headers for document generation
huffman
parents: 24333
diff changeset
     6
29628
d9294387ab0e entry point for Word library now named Word
haftmann
parents: 27137
diff changeset
     7
theory Word
41413
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 41060
diff changeset
     8
imports
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 65363
diff changeset
     9
  "HOL-Library.Type_Length"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
    10
begin
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
    11
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    12
subsection \<open>Preliminaries\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    13
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    14
lemma signed_take_bit_decr_length_iff:
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    15
  \<open>signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    16
    \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
    17
  by (simp add: signed_take_bit_eq_iff_take_bit_eq)
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    18
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    19
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    20
subsection \<open>Fundamentals\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    21
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    22
subsubsection \<open>Type definition\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
    23
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
    24
quotient_type (overloaded) 'a word = int / \<open>\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    25
  morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI)
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    26
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    27
hide_const (open) rep \<comment> \<open>only for foundational purpose\<close>
72130
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
    28
hide_const (open) Word \<comment> \<open>only for code generation\<close>
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
    29
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    30
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    31
subsubsection \<open>Basic arithmetic\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    32
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    33
instantiation word :: (len) comm_ring_1
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    34
begin
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    35
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    36
lift_definition zero_word :: \<open>'a word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    37
  is 0 .
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    38
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    39
lift_definition one_word :: \<open>'a word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    40
  is 1 .
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    41
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    42
lift_definition plus_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    43
  is \<open>(+)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
    44
  by (auto simp: take_bit_eq_mod intro: mod_add_cong)
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    45
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    46
lift_definition minus_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    47
  is \<open>(-)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
    48
  by (auto simp: take_bit_eq_mod intro: mod_diff_cong)
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    49
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    50
lift_definition uminus_word :: \<open>'a word \<Rightarrow> 'a word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    51
  is uminus
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
    52
  by (auto simp: take_bit_eq_mod intro: mod_minus_cong)
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    53
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    54
lift_definition times_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    55
  is \<open>(*)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
    56
  by (auto simp: take_bit_eq_mod intro: mod_mult_cong)
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    57
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    58
instance
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    59
  by (standard; transfer) (simp_all add: algebra_simps)
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    60
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    61
end
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    62
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    63
context
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    64
  includes lifting_syntax
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    65
  notes
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    66
    power_transfer [transfer_rule]
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    67
    transfer_rule_of_bool [transfer_rule]
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    68
    transfer_rule_numeral [transfer_rule]
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    69
    transfer_rule_of_nat [transfer_rule]
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    70
    transfer_rule_of_int [transfer_rule]
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    71
begin
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    72
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    73
lemma power_transfer_word [transfer_rule]:
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    74
  \<open>(pcr_word ===> (=) ===> pcr_word) (^) (^)\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    75
  by transfer_prover
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
    76
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    77
lemma [transfer_rule]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    78
  \<open>((=) ===> pcr_word) of_bool of_bool\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    79
  by transfer_prover
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    80
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    81
lemma [transfer_rule]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    82
  \<open>((=) ===> pcr_word) numeral numeral\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    83
  by transfer_prover
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    84
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    85
lemma [transfer_rule]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    86
  \<open>((=) ===> pcr_word) int of_nat\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    87
  by transfer_prover
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    88
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    89
lemma [transfer_rule]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    90
  \<open>((=) ===> pcr_word) (\<lambda>k. k) of_int\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    91
proof -
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    92
  have \<open>((=) ===> pcr_word) of_int of_int\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    93
    by transfer_prover
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    94
  then show ?thesis by (simp add: id_def)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    95
qed
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    96
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    97
lemma [transfer_rule]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    98
  \<open>(pcr_word ===> (\<longleftrightarrow>)) even ((dvd) 2 :: 'a::len word \<Rightarrow> bool)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
    99
proof -
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   100
  have even_word_unfold: "even k \<longleftrightarrow> (\<exists>l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \<longleftrightarrow> ?Q")
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   101
    for k :: int
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   102
    by (metis dvd_triv_left evenE even_take_bit_eq len_not_eq_0)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   103
  show ?thesis 
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   104
    unfolding even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   105
    by transfer_prover
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   106
qed
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   107
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   108
end
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   109
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
   110
lemma exp_eq_zero_iff [simp]:
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
   111
  \<open>2 ^ n = (0 :: 'a::len word) \<longleftrightarrow> n \<ge> LENGTH('a)\<close>
73535
0f33c7031ec9 new lemmas
haftmann
parents: 72954
diff changeset
   112
  by transfer auto
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
   113
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   114
lemma word_exp_length_eq_0 [simp]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   115
  \<open>(2 :: 'a::len word) ^ LENGTH('a) = 0\<close>
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
   116
  by simp
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   117
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   118
72489
a1366ce41368 early and more complete setup of tools
haftmann
parents: 72488
diff changeset
   119
subsubsection \<open>Basic tool setup\<close>
a1366ce41368 early and more complete setup of tools
haftmann
parents: 72488
diff changeset
   120
a1366ce41368 early and more complete setup of tools
haftmann
parents: 72488
diff changeset
   121
ML_file \<open>Tools/word_lib.ML\<close>
a1366ce41368 early and more complete setup of tools
haftmann
parents: 72488
diff changeset
   122
a1366ce41368 early and more complete setup of tools
haftmann
parents: 72488
diff changeset
   123
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   124
subsubsection \<open>Basic code generation setup\<close>
71948
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
   125
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   126
context
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   127
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   128
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   129
qualified lift_definition the_int :: \<open>'a::len word \<Rightarrow> int\<close>
71948
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
   130
  is \<open>take_bit LENGTH('a)\<close> .
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
   131
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   132
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   133
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   134
lemma [code abstype]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   135
  \<open>Word.Word (Word.the_int w) = w\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   136
  by transfer simp
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   137
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   138
lemma Word_eq_word_of_int [code_post, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   139
  \<open>Word.Word = of_int\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   140
  by (rule; transfer) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   141
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   142
quickcheck_generator word
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   143
  constructors:
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   144
    \<open>0 :: 'a::len word\<close>,
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   145
    \<open>numeral :: num \<Rightarrow> 'a::len word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   146
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   147
instantiation word :: (len) equal
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   148
begin
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   149
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   150
lift_definition equal_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> bool\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   151
  is \<open>\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   152
  by simp
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   153
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   154
instance
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   155
  by (standard; transfer) rule
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   156
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   157
end
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   158
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   159
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   160
  \<open>HOL.equal v w \<longleftrightarrow> HOL.equal (Word.the_int v) (Word.the_int w)\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   161
  by transfer (simp add: equal)
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   162
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   163
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   164
  \<open>Word.the_int 0 = 0\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   165
  by transfer simp
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   166
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   167
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   168
  \<open>Word.the_int 1 = 1\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   169
  by transfer simp
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   170
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   171
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   172
  \<open>Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   173
  for v w :: \<open>'a::len word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   174
  by transfer (simp add: take_bit_add)
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   175
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   176
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   177
  \<open>Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   178
  for w :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   179
  by transfer (auto simp: take_bit_eq_mod zmod_zminus1_eq_if)
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   180
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   181
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   182
  \<open>Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   183
  for v w :: \<open>'a::len word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   184
  by transfer (simp add: take_bit_diff)
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   185
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   186
lemma [code]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   187
  \<open>Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\<close>
72243
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   188
  for v w :: \<open>'a::len word\<close>
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   189
  by transfer (simp add: take_bit_mult)
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   190
eaac77208cf9 tuned theory structure
haftmann
parents: 72242
diff changeset
   191
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   192
subsubsection \<open>Basic conversions\<close>
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
   193
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   194
abbreviation word_of_nat :: \<open>nat \<Rightarrow> 'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   195
  where \<open>word_of_nat \<equiv> of_nat\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   196
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   197
abbreviation word_of_int :: \<open>int \<Rightarrow> 'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   198
  where \<open>word_of_int \<equiv> of_int\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   199
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   200
lemma word_of_nat_eq_iff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   201
  \<open>word_of_nat m = (word_of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m = take_bit LENGTH('a) n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   202
  by transfer (simp add: take_bit_of_nat)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   203
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   204
lemma word_of_int_eq_iff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   205
  \<open>word_of_int k = (word_of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   206
  by transfer rule
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   207
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
   208
lemma word_of_nat_eq_0_iff:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   209
  \<open>word_of_nat n = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   210
  using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   211
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
   212
lemma word_of_int_eq_0_iff:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   213
  \<open>word_of_int k = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd k\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   214
  using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   215
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   216
context semiring_1
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   217
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   218
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   219
lift_definition unsigned :: \<open>'b::len word \<Rightarrow> 'a\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   220
  is \<open>of_nat \<circ> nat \<circ> take_bit LENGTH('b)\<close>
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   221
  by simp
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
   222
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   223
lemma unsigned_0 [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   224
  \<open>unsigned 0 = 0\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   225
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   226
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   227
lemma unsigned_1 [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   228
  \<open>unsigned 1 = 1\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   229
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   230
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   231
lemma unsigned_numeral [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   232
  \<open>unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   233
  by transfer (simp add: nat_take_bit_eq)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   234
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   235
lemma unsigned_neg_numeral [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   236
  \<open>unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   237
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   238
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   239
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   240
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   241
context semiring_1
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   242
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   243
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
   244
lemma unsigned_of_nat:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   245
  \<open>unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   246
  by transfer (simp add: nat_eq_iff take_bit_of_nat)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   247
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
   248
lemma unsigned_of_int:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   249
  \<open>unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   250
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   251
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   252
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   253
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   254
context semiring_char_0
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   255
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   256
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   257
lemma unsigned_word_eqI:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   258
  \<open>v = w\<close> if \<open>unsigned v = unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   259
  using that by transfer (simp add: eq_nat_nat_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   260
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   261
lemma word_eq_iff_unsigned:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   262
  \<open>v = w \<longleftrightarrow> unsigned v = unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   263
  by (auto intro: unsigned_word_eqI)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   264
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   265
lemma inj_unsigned [simp]:
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   266
  \<open>inj unsigned\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   267
  by (rule injI) (simp add: unsigned_word_eqI)
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   268
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   269
lemma unsigned_eq_0_iff:
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   270
  \<open>unsigned w = 0 \<longleftrightarrow> w = 0\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   271
  using word_eq_iff_unsigned [of w 0] by simp
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   272
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   273
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   274
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   275
context ring_1
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   276
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   277
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   278
lift_definition signed :: \<open>'b::len word \<Rightarrow> 'a\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   279
  is \<open>of_int \<circ> signed_take_bit (LENGTH('b) - Suc 0)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   280
  by (simp flip: signed_take_bit_decr_length_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   281
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   282
lemma signed_0 [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   283
  \<open>signed 0 = 0\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   284
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   285
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   286
lemma signed_1 [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   287
  \<open>signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
   288
  by (transfer fixing: uminus; cases \<open>LENGTH('b)\<close>) (auto dest: gr0_implies_Suc)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   289
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   290
lemma signed_minus_1 [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   291
  \<open>signed (- 1 :: 'b::len word) = - 1\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   292
  by (transfer fixing: uminus) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   293
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   294
lemma signed_numeral [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   295
  \<open>signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   296
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   297
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   298
lemma signed_neg_numeral [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   299
  \<open>signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   300
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   301
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
   302
lemma signed_of_nat:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   303
  \<open>signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   304
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   305
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
   306
lemma signed_of_int:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   307
  \<open>signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   308
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   309
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   310
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   311
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   312
context ring_char_0
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   313
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   314
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   315
lemma signed_word_eqI:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   316
  \<open>v = w\<close> if \<open>signed v = signed w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   317
  using that by transfer (simp flip: signed_take_bit_decr_length_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   318
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   319
lemma word_eq_iff_signed:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   320
  \<open>v = w \<longleftrightarrow> signed v = signed w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   321
  by (auto intro: signed_word_eqI)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   322
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   323
lemma inj_signed [simp]:
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   324
  \<open>inj signed\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   325
  by (rule injI) (simp add: signed_word_eqI)
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
   326
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   327
lemma signed_eq_0_iff:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   328
  \<open>signed w = 0 \<longleftrightarrow> w = 0\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   329
  using word_eq_iff_signed [of w 0] by simp
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   330
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   331
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   332
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   333
abbreviation unat :: \<open>'a::len word \<Rightarrow> nat\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   334
  where \<open>unat \<equiv> unsigned\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   335
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   336
abbreviation uint :: \<open>'a::len word \<Rightarrow> int\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   337
  where \<open>uint \<equiv> unsigned\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   338
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   339
abbreviation sint :: \<open>'a::len word \<Rightarrow> int\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   340
  where \<open>sint \<equiv> signed\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   341
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   342
abbreviation ucast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   343
  where \<open>ucast \<equiv> unsigned\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   344
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   345
abbreviation scast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   346
  where \<open>scast \<equiv> signed\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   347
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   348
context
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   349
  includes lifting_syntax
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   350
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   351
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   352
lemma [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   353
  \<open>(pcr_word ===> (=)) (nat \<circ> take_bit LENGTH('a)) (unat :: 'a::len word \<Rightarrow> nat)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   354
  using unsigned.transfer [where ?'a = nat] by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   355
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   356
lemma [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   357
  \<open>(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \<Rightarrow> int)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   358
  using unsigned.transfer [where ?'a = int] by (simp add: comp_def)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   359
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   360
lemma [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   361
  \<open>(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \<Rightarrow> int)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   362
  using signed.transfer [where ?'a = int] by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   363
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   364
lemma [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   365
  \<open>(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \<Rightarrow> 'b::len word)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   366
proof (rule rel_funI)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   367
  fix k :: int and w :: \<open>'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   368
  assume \<open>pcr_word k w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   369
  then have \<open>w = word_of_int k\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   370
    by (simp add: pcr_word_def cr_word_def relcompp_apply)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   371
  moreover have \<open>pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   372
    by transfer (simp add: pcr_word_def cr_word_def relcompp_apply)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   373
  ultimately show \<open>pcr_word (take_bit LENGTH('a) k) (ucast w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   374
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   375
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   376
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   377
lemma [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   378
  \<open>(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \<Rightarrow> 'b::len word)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   379
proof (rule rel_funI)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   380
  fix k :: int and w :: \<open>'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   381
  assume \<open>pcr_word k w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   382
  then have \<open>w = word_of_int k\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   383
    by (simp add: pcr_word_def cr_word_def relcompp_apply)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   384
  moreover have \<open>pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   385
    by transfer (simp add: pcr_word_def cr_word_def relcompp_apply)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   386
  ultimately show \<open>pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   387
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   388
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   389
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   390
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   391
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   392
lemma of_nat_unat [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   393
  \<open>of_nat (unat w) = unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   394
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   395
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   396
lemma of_int_uint [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   397
  \<open>of_int (uint w) = unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   398
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   399
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   400
lemma of_int_sint [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   401
  \<open>of_int (sint a) = signed a\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   402
  by transfer (simp_all add: take_bit_signed_take_bit)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
   403
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
   404
lemma nat_uint_eq [simp]:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
   405
  \<open>nat (uint w) = unat w\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
   406
  by transfer simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
   407
81609
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   408
lemma nat_of_natural_unsigned_eq [simp]:
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   409
  \<open>nat_of_natural (unsigned w) = unat w\<close>
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   410
  by transfer simp
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   411
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   412
lemma int_of_integer_unsigned_eq [simp]:
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   413
  \<open>int_of_integer (unsigned w) = uint w\<close>
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   414
  by transfer simp
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   415
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   416
lemma int_of_integer_signed_eq [simp]:
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   417
  \<open>int_of_integer (signed w) = sint w\<close>
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   418
  by transfer simp
3458f17e7cba more simp rules on word conversions
haftmann
parents: 81142
diff changeset
   419
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   420
lemma sgn_uint_eq [simp]:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   421
  \<open>sgn (uint w) = of_bool (w \<noteq> 0)\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   422
  by transfer (simp add: less_le)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
   423
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   424
text \<open>Aliasses only for code generation\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   425
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   426
context
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   427
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   428
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   429
qualified lift_definition of_int :: \<open>int \<Rightarrow> 'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   430
  is \<open>take_bit LENGTH('a)\<close> .
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   431
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   432
qualified lift_definition of_nat :: \<open>nat \<Rightarrow> 'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   433
  is \<open>int \<circ> take_bit LENGTH('a)\<close> .
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   434
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   435
qualified lift_definition the_nat :: \<open>'a::len word \<Rightarrow> nat\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   436
  is \<open>nat \<circ> take_bit LENGTH('a)\<close> by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   437
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   438
qualified lift_definition the_signed_int :: \<open>'a::len word \<Rightarrow> int\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   439
  is \<open>signed_take_bit (LENGTH('a) - Suc 0)\<close> by (simp add: signed_take_bit_decr_length_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   440
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   441
qualified lift_definition cast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   442
  is \<open>take_bit LENGTH('a)\<close> by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   443
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   444
qualified lift_definition signed_cast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   445
  is \<open>signed_take_bit (LENGTH('a) - Suc 0)\<close> by (metis signed_take_bit_decr_length_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   446
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   447
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   448
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   449
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   450
  \<open>Word.the_int = uint\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   451
  by transfer rule
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   452
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   453
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   454
  \<open>Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   455
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   456
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   457
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   458
  \<open>Word.of_int = word_of_int\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   459
  by (rule; transfer) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   460
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   461
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   462
  \<open>Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\<close>
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   463
  by transfer (simp add: take_bit_of_nat)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   464
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   465
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   466
  \<open>Word.of_nat = word_of_nat\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   467
  by (rule; transfer) (simp add: take_bit_of_nat)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   468
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   469
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   470
  \<open>Word.the_nat w = nat (Word.the_int w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   471
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   472
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   473
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   474
  \<open>Word.the_nat = unat\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   475
  by (rule; transfer) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   476
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   477
lemma [code]:
82367
07e95760a612 simplified implementation of the_signed_int
haftmann
parents: 81763
diff changeset
   478
  \<open>Word.the_signed_int w = (let k = Word.the_int w
07e95760a612 simplified implementation of the_signed_int
haftmann
parents: 81763
diff changeset
   479
    in if bit k (LENGTH('a) - Suc 0) then k + push_bit LENGTH('a) (- 1) else k)\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   480
  for w :: \<open>'a::len word\<close>
82367
07e95760a612 simplified implementation of the_signed_int
haftmann
parents: 81763
diff changeset
   481
  by transfer (simp add: bit_simps signed_take_bit_eq_take_bit_add)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   482
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   483
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   484
  \<open>Word.the_signed_int = sint\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   485
  by (rule; transfer) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   486
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   487
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   488
  \<open>Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   489
  for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   490
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   491
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   492
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   493
  \<open>Word.cast = ucast\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   494
  by (rule; transfer) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   495
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   496
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   497
  \<open>Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   498
  for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   499
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   500
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   501
lemma [code_abbrev, simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   502
  \<open>Word.signed_cast = scast\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   503
  by (rule; transfer) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   504
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   505
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   506
  \<open>unsigned w = of_nat (nat (Word.the_int w))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   507
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   508
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   509
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   510
  \<open>signed w = of_int (Word.the_signed_int w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   511
  by transfer simp
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   512
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   513
82526
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   514
subsection \<open>Elementary case distinctions\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   515
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   516
lemma word_length_one [case_names zero minus_one length_beyond]:
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   517
  fixes w :: \<open>'a::len word\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   518
  obtains (zero) \<open>LENGTH('a) = Suc 0\<close> \<open>w = 0\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   519
  | (minus_one) \<open>LENGTH('a) = Suc 0\<close> \<open>w = - 1\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   520
  | (length_beyond) \<open>2 \<le> LENGTH('a)\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   521
proof (cases \<open>2 \<le> LENGTH('a)\<close>)
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   522
  case True
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   523
  with length_beyond show ?thesis .
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   524
next
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   525
  case False
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   526
  then have \<open>LENGTH('a) = Suc 0\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   527
    by simp
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   528
  then have \<open>w = 0 \<or> w = - 1\<close>
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   529
    by transfer auto
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   530
  with \<open>LENGTH('a) = Suc 0\<close> zero minus_one show ?thesis
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   531
    by blast
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   532
qed
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   533
c49564b6eb0f explicit case rule
haftmann
parents: 82525
diff changeset
   534
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   535
subsubsection \<open>Basic ordering\<close>
45547
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   536
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
   537
instantiation word :: (len) linorder
45547
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   538
begin
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   539
71950
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   540
lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   541
  is "\<lambda>a b. take_bit LENGTH('a) a \<le> take_bit LENGTH('a) b"
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   542
  by simp
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   543
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   544
lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   545
  is "\<lambda>a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b"
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   546
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
   547
45547
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   548
instance
71950
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   549
  by (standard; transfer) auto
45547
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   550
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   551
end
94c37f3df10f HOL-Word: removed more duplicate theorems
huffman
parents: 45546
diff changeset
   552
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
   553
interpretation word_order: ordering_top \<open>(\<le>)\<close> \<open>(<)\<close> \<open>- 1 :: 'a::len word\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
   554
  by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1)
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
   555
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
   556
interpretation word_coorder: ordering_top \<open>(\<ge>)\<close> \<open>(>)\<close> \<open>0 :: 'a::len word\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
   557
  by (standard; transfer) simp
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
   558
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   559
lemma word_of_nat_less_eq_iff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   560
  \<open>word_of_nat m \<le> (word_of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   561
  by transfer (simp add: take_bit_of_nat)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   562
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   563
lemma word_of_int_less_eq_iff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   564
  \<open>word_of_int k \<le> (word_of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   565
  by transfer rule
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   566
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   567
lemma word_of_nat_less_iff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   568
  \<open>word_of_nat m < (word_of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   569
  by transfer (simp add: take_bit_of_nat)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   570
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   571
lemma word_of_int_less_iff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   572
  \<open>word_of_int k < (word_of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   573
  by transfer rule
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   574
71950
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   575
lemma word_le_def [code]:
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   576
  "a \<le> b \<longleftrightarrow> uint a \<le> uint b"
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   577
  by transfer rule
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   578
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   579
lemma word_less_def [code]:
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   580
  "a < b \<longleftrightarrow> uint a < uint b"
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   581
  by transfer rule
c9251bc7da4e more transfer rules
haftmann
parents: 71949
diff changeset
   582
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   583
lemma word_greater_zero_iff:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
   584
  \<open>a > 0 \<longleftrightarrow> a \<noteq> 0\<close> for a :: \<open>'a::len word\<close>
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   585
  by transfer (simp add: less_le)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   586
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   587
lemma of_nat_word_less_eq_iff:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   588
  \<open>of_nat m \<le> (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   589
  by transfer (simp add: take_bit_of_nat)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   590
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   591
lemma of_nat_word_less_iff:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   592
  \<open>of_nat m < (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   593
  by transfer (simp add: take_bit_of_nat)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   594
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   595
lemma of_int_word_less_eq_iff:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   596
  \<open>of_int k \<le> (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   597
  by transfer rule
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   598
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   599
lemma of_int_word_less_iff:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   600
  \<open>of_int k < (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   601
  by transfer rule
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   602
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
   603
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   604
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   605
subsection \<open>Enumeration\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   606
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   607
lemma inj_on_word_of_nat:
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   608
  \<open>inj_on (word_of_nat :: nat \<Rightarrow> 'a::len word) {0..<2 ^ LENGTH('a)}\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   609
  by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self)
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   610
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   611
lemma UNIV_word_eq_word_of_nat:
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   612
  \<open>(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\<close> (is \<open>_ = ?A\<close>)
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   613
proof
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   614
  show \<open>word_of_nat ` {0..<2 ^ LENGTH('a)} \<subseteq> UNIV\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   615
    by simp
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   616
  show \<open>UNIV \<subseteq> ?A\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   617
  proof
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   618
    fix w :: \<open>'a word\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   619
    show \<open>w \<in> (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   620
      by (rule image_eqI [of _ _ \<open>unat w\<close>]; transfer) simp_all
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   621
  qed
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   622
qed
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   623
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   624
instantiation word :: (len) enum
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   625
begin
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   626
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   627
definition enum_word :: \<open>'a word list\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   628
  where \<open>enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   629
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   630
definition enum_all_word :: \<open>('a word \<Rightarrow> bool) \<Rightarrow> bool\<close>
77225
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   631
  where \<open>enum_all_word = All\<close>
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   632
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   633
definition enum_ex_word :: \<open>('a word \<Rightarrow> bool) \<Rightarrow> bool\<close>
77225
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   634
  where \<open>enum_ex_word = Ex\<close>
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   635
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   636
instance
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   637
  by standard
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   638
    (simp_all add: enum_all_word_def enum_ex_word_def enum_word_def distinct_map inj_on_word_of_nat flip: UNIV_word_eq_word_of_nat)
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   639
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   640
end
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   641
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   642
lemma [code]:
77225
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   643
  \<open>Enum.enum_all P \<longleftrightarrow> list_all P Enum.enum\<close>
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   644
  \<open>Enum.enum_ex P \<longleftrightarrow> list_ex P Enum.enum\<close> for P :: \<open>'a::len word \<Rightarrow> bool\<close>
b6f3eb537d91 actually executable enum_all, enum_ex for word
haftmann
parents: 77061
diff changeset
   645
  by (simp_all add: enum_all_word_def enum_ex_word_def enum_UNIV list_all_iff list_ex_iff)
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   646
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
   647
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
   648
subsection \<open>Bit-wise operations\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
   649
77812
fb3d81bd9803 some remarks on division
haftmann
parents: 77225
diff changeset
   650
text \<open>
fb3d81bd9803 some remarks on division
haftmann
parents: 77225
diff changeset
   651
  The following specification of word division just lifts the pre-existing
79950
82aaa0d8fc3b isabelle update -u cite;
wenzelm
parents: 79673
diff changeset
   652
  division on integers named ``F-Division'' in \<^cite>\<open>"leijen01"\<close>.
77812
fb3d81bd9803 some remarks on division
haftmann
parents: 77225
diff changeset
   653
\<close>
fb3d81bd9803 some remarks on division
haftmann
parents: 77225
diff changeset
   654
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   655
instantiation word :: (len) semiring_modulo
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   656
begin
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   657
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   658
lift_definition divide_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   659
  is \<open>\<lambda>a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   660
  by simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   661
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   662
lift_definition modulo_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   663
  is \<open>\<lambda>a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   664
  by simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   665
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   666
instance proof
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   667
  show "a div b * b + a mod b = a" for a b :: "'a word"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   668
  proof transfer
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   669
    fix k l :: int
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   670
    define r :: int where "r = 2 ^ LENGTH('a)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   671
    then have r: "take_bit LENGTH('a) k = k mod r" for k
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   672
      by (simp add: take_bit_eq_mod)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   673
    have "k mod r = ((k mod r) div (l mod r) * (l mod r)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   674
      + (k mod r) mod (l mod r)) mod r"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   675
      by (simp add: div_mult_mod_eq)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   676
    also have "\<dots> = (((k mod r) div (l mod r) * (l mod r)) mod r
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   677
      + (k mod r) mod (l mod r)) mod r"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   678
      by (simp add: mod_add_left_eq)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   679
    also have "\<dots> = (((k mod r) div (l mod r) * l) mod r
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   680
      + (k mod r) mod (l mod r)) mod r"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   681
      by (simp add: mod_mult_right_eq)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   682
    finally have "k mod r = ((k mod r) div (l mod r) * l
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   683
      + (k mod r) mod (l mod r)) mod r"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   684
      by (simp add: mod_simps)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   685
    with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   686
      + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   687
      by simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   688
  qed
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   689
qed
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   690
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   691
end
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   692
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   693
lemma unat_div_distrib:
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   694
  \<open>unat (v div w) = unat v div unat w\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   695
proof transfer
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   696
  fix k l
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   697
  have \<open>nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \<le> nat (take_bit LENGTH('a) k)\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   698
    by (rule div_le_dividend)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   699
  also have \<open>nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   700
    by (simp add: nat_less_iff)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   701
  finally show \<open>(nat \<circ> take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) =
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   702
    (nat \<circ> take_bit LENGTH('a)) k div (nat \<circ> take_bit LENGTH('a)) l\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   703
    by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   704
qed
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   705
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   706
lemma unat_mod_distrib:
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   707
  \<open>unat (v mod w) = unat v mod unat w\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   708
proof transfer
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   709
  fix k l
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   710
  have \<open>nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \<le> nat (take_bit LENGTH('a) k)\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   711
    by (rule mod_less_eq_dividend)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   712
  also have \<open>nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   713
    by (simp add: nat_less_iff)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   714
  finally show \<open>(nat \<circ> take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) =
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   715
    (nat \<circ> take_bit LENGTH('a)) k mod (nat \<circ> take_bit LENGTH('a)) l\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   716
    by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   717
qed
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   718
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   719
instance word :: (len) semiring_parity
79555
8ef205d9fd22 strengthened class parity
haftmann
parents: 79531
diff changeset
   720
  by (standard; transfer)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   721
    (auto simp: mod_2_eq_odd take_bit_Suc elim: evenE dest: le_Suc_ex)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
   722
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   723
lemma word_bit_induct [case_names zero even odd]:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   724
  \<open>P a\<close> if word_zero: \<open>P 0\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   725
    and word_even: \<open>\<And>a. P a \<Longrightarrow> 0 < a \<Longrightarrow> a < 2 ^ (LENGTH('a) - Suc 0) \<Longrightarrow> P (2 * a)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   726
    and word_odd: \<open>\<And>a. P a \<Longrightarrow> a < 2 ^ (LENGTH('a) - Suc 0) \<Longrightarrow> P (1 + 2 * a)\<close>
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   727
  for P and a :: \<open>'a::len word\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   728
proof -
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   729
  define m :: nat where \<open>m = LENGTH('a) - Suc 0\<close>
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   730
  then have l: \<open>LENGTH('a) = Suc m\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   731
    by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   732
  define n :: nat where \<open>n = unat a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   733
  then have \<open>n < 2 ^ LENGTH('a)\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   734
    by transfer (simp add: take_bit_eq_mod)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   735
  then have \<open>n < 2 * 2 ^ m\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   736
    by (simp add: l)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   737
  then have \<open>P (of_nat n)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   738
  proof (induction n rule: nat_bit_induct)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   739
    case zero
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   740
    show ?case
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   741
      by simp (rule word_zero)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   742
  next
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   743
    case (even n)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   744
    then have \<open>n < 2 ^ m\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   745
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   746
    with even.IH have \<open>P (of_nat n)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   747
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   748
    moreover from \<open>n < 2 ^ m\<close> even.hyps have \<open>0 < (of_nat n :: 'a word)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   749
      by (auto simp: word_greater_zero_iff l word_of_nat_eq_0_iff)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   750
    moreover from \<open>n < 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\<close>
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   751
      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
72261
5193570b739a more lemmas
haftmann
parents: 72244
diff changeset
   752
      by (simp add: l take_bit_eq_mod)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   753
    ultimately have \<open>P (2 * of_nat n)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   754
      by (rule word_even)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   755
    then show ?case
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   756
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   757
  next
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   758
    case (odd n)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   759
    then have \<open>Suc n \<le> 2 ^ m\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   760
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   761
    with odd.IH have \<open>P (of_nat n)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   762
      by simp
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   763
    moreover from \<open>Suc n \<le> 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\<close>
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   764
      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
72261
5193570b739a more lemmas
haftmann
parents: 72244
diff changeset
   765
      by (simp add: l take_bit_eq_mod)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   766
    ultimately have \<open>P (1 + 2 * of_nat n)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   767
      by (rule word_odd)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   768
    then show ?case
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   769
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   770
  qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   771
  moreover have \<open>of_nat (nat (uint a)) = a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   772
    by transfer simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   773
  ultimately show ?thesis
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
   774
    by (simp add: n_def)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   775
qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   776
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   777
lemma bit_word_half_eq:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   778
  \<open>(of_bool b + a * 2) div 2 = a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   779
    if \<open>a < 2 ^ (LENGTH('a) - Suc 0)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   780
    for a :: \<open>'a::len word\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   781
proof (cases \<open>2 \<le> LENGTH('a::len)\<close>)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   782
  case False
82524
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
   783
  with that show ?thesis
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
   784
    by transfer simp
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   785
next
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   786
  case True
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   787
  obtain n where length: \<open>LENGTH('a) = Suc n\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   788
    by (cases \<open>LENGTH('a)\<close>) simp_all
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   789
  show ?thesis proof (cases b)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   790
    case False
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   791
    moreover have \<open>a * 2 div 2 = a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   792
    using that proof transfer
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   793
      fix k :: int
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   794
      from length have \<open>k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   795
        by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   796
      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 75623
diff changeset
   797
      with \<open>LENGTH('a) = Suc n\<close> have \<open>take_bit LENGTH('a) k = take_bit n k\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   798
        by (auto simp: take_bit_Suc_from_most)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   799
      ultimately have \<open>take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   800
        by (simp add: take_bit_eq_mod)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   801
      with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   802
        = take_bit LENGTH('a) k\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   803
        by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   804
    qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   805
    ultimately show ?thesis
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   806
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   807
  next
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   808
    case True
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   809
    moreover have \<open>(1 + a * 2) div 2 = a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   810
    using that proof transfer
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   811
      fix k :: int
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   812
      from length have \<open>(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   813
        using pos_zmod_mult_2 [of \<open>2 ^ n\<close> k] by (simp add: ac_simps)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   814
      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 75623
diff changeset
   815
      with \<open>LENGTH('a) = Suc n\<close> have \<open>take_bit LENGTH('a) k = take_bit n k\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   816
        by (auto simp: take_bit_Suc_from_most)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   817
      ultimately have \<open>take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   818
        by (simp add: take_bit_eq_mod)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   819
      with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   820
        = take_bit LENGTH('a) k\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   821
        by (auto simp: take_bit_Suc)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   822
    qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   823
    ultimately show ?thesis
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   824
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   825
  qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   826
qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   827
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   828
lemma even_mult_exp_div_word_iff:
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   829
  \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> \<not> (
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   830
    m \<le> n \<and>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   831
    n < LENGTH('a) \<and> odd (a div 2 ^ (n - m)))\<close> for a :: \<open>'a::len word\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   832
  by transfer
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   833
    (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff,
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   834
      simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   835
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   836
instantiation word :: (len) semiring_bits
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   837
begin
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   838
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   839
lift_definition bit_word :: \<open>'a word \<Rightarrow> nat \<Rightarrow> bool\<close>
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   840
  is \<open>\<lambda>k n. n < LENGTH('a) \<and> bit k n\<close>
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   841
proof
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   842
  fix k l :: int and n :: nat
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   843
  assume *: \<open>take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   844
  show \<open>n < LENGTH('a) \<and> bit k n \<longleftrightarrow> n < LENGTH('a) \<and> bit l n\<close>
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   845
  proof (cases \<open>n < LENGTH('a)\<close>)
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   846
    case True
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   847
    from * have \<open>bit (take_bit LENGTH('a) k) n \<longleftrightarrow> bit (take_bit LENGTH('a) l) n\<close>
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   848
      by simp
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   849
    then show ?thesis
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   850
      by (simp add: bit_take_bit_iff)
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   851
  next
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   852
    case False
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   853
    then show ?thesis
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   854
      by simp
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   855
  qed
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   856
qed
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   857
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   858
instance proof
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   859
  show \<open>P a\<close> if stable: \<open>\<And>a. a div 2 = a \<Longrightarrow> P a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   860
    and rec: \<open>\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   861
  for P and a :: \<open>'a word\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   862
  proof (induction a rule: word_bit_induct)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   863
    case zero
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   864
    have \<open>0 div 2 = (0::'a word)\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   865
      by transfer simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   866
    with stable [of 0] show ?case
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   867
      by simp
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   868
  next
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   869
    case (even a)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   870
    with rec [of a False] show ?case
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   871
      using bit_word_half_eq [of a False] by (simp add: ac_simps)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   872
  next
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   873
    case (odd a)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   874
    with rec [of a True] show ?case
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   875
      using bit_word_half_eq [of a True] by (simp add: ac_simps)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   876
  qed
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   877
  show \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close> for a :: \<open>'a word\<close> and n
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71958
diff changeset
   878
    by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit)
79481
8205977e9e2c simplified specification of type class
haftmann
parents: 79118
diff changeset
   879
  show \<open>a div 0 = 0\<close>
8205977e9e2c simplified specification of type class
haftmann
parents: 79118
diff changeset
   880
    for a :: \<open>'a word\<close>
8205977e9e2c simplified specification of type class
haftmann
parents: 79118
diff changeset
   881
    by transfer simp
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   882
  show \<open>a div 1 = a\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   883
    for a :: \<open>'a word\<close>
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   884
    by transfer simp
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   885
  show \<open>0 div a = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   886
    for a :: \<open>'a word\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   887
    by transfer simp
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   888
  show \<open>a mod b div b = 0\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   889
    for a b :: \<open>'a word\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79590
diff changeset
   890
    by (simp add: word_eq_iff_unsigned [where ?'a = nat] unat_div_distrib unat_mod_distrib)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   891
  show \<open>a div 2 div 2 ^ n = a div 2 ^ Suc n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   892
    for a :: \<open>'a word\<close> and m n :: nat
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   893
    apply transfer
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   894
    using drop_bit_eq_div [symmetric, where ?'a = int,of _ 1]
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   895
    apply (auto simp:  not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div simp del: power.simps)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   896
    apply (simp add: drop_bit_take_bit)
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   897
    done
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   898
  show \<open>even (2 * a div 2 ^ Suc n) \<longleftrightarrow> even (a div 2 ^ n)\<close> if \<open>2 ^ Suc n \<noteq> (0::'a word)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   899
    for a :: \<open>'a word\<close> and n :: nat
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   900
    using that by transfer
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   901
      (simp add: even_drop_bit_iff_not_bit bit_simps flip: drop_bit_eq_div del: power.simps)
71951
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   902
qed
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   903
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   904
end
ac6f9738c200 essential instance about bit structure
haftmann
parents: 71950
diff changeset
   905
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   906
lemma bit_word_eqI:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   907
  \<open>a = b\<close> if \<open>\<And>n. n < LENGTH('a) \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   908
  for a b :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   909
  using that by transfer (auto simp: nat_less_le bit_eq_iff bit_take_bit_iff)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   910
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   911
lemma bit_imp_le_length: \<open>n < LENGTH('a)\<close> if \<open>bit w n\<close> for w :: \<open>'a::len word\<close>
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   912
  by (meson bit_word.rep_eq that)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   913
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   914
lemma not_bit_length [simp]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   915
  \<open>\<not> bit w LENGTH('a)\<close> for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   916
  using bit_imp_le_length by blast
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   917
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72735
diff changeset
   918
lemma finite_bit_word [simp]:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72735
diff changeset
   919
  \<open>finite {n. bit w n}\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72735
diff changeset
   920
  for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   921
  by (metis bit_imp_le_length bounded_nat_set_is_finite mem_Collect_eq)
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72735
diff changeset
   922
73789
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   923
lemma bit_numeral_word_iff [simp]:
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   924
  \<open>bit (numeral w :: 'a::len word) n
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   925
    \<longleftrightarrow> n < LENGTH('a) \<and> bit (numeral w :: int) n\<close>
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   926
  by transfer simp
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   927
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   928
lemma bit_neg_numeral_word_iff [simp]:
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   929
  \<open>bit (- numeral w :: 'a::len word) n
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   930
    \<longleftrightarrow> n < LENGTH('a) \<and> bit (- numeral w :: int) n\<close>
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   931
  by transfer simp
aab7975fa070 more lemmas
haftmann
parents: 73788
diff changeset
   932
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   933
instantiation word :: (len) ring_bit_operations
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   934
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   935
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   936
lift_definition not_word :: \<open>'a word \<Rightarrow> 'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   937
  is not
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   938
  by (simp add: take_bit_not_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   939
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   940
lift_definition and_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   941
  is \<open>and\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   942
  by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   943
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   944
lift_definition or_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   945
  is or
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   946
  by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   947
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   948
lift_definition xor_word ::  \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   949
  is xor
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   950
  by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   951
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   952
lift_definition mask_word :: \<open>nat \<Rightarrow> 'a word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   953
  is mask
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   954
  .
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   955
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   956
lift_definition set_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   957
  is set_bit
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   958
  by (simp add: set_bit_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   959
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   960
lift_definition unset_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   961
  is unset_bit
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   962
  by (simp add: unset_bit_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   963
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   964
lift_definition flip_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   965
  is flip_bit
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   966
  by (simp add: flip_bit_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   967
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   968
lift_definition push_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   969
  is push_bit
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   970
proof -
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   971
  show \<open>take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   972
    if \<open>take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close> for k l :: int and n :: nat
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
   973
    by (metis le_add2 push_bit_take_bit take_bit_tightened that)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   974
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   975
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   976
lift_definition drop_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   977
  is \<open>\<lambda>n. drop_bit n \<circ> take_bit LENGTH('a)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   978
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   979
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   980
lift_definition take_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   981
  is \<open>\<lambda>n. take_bit (min LENGTH('a) n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   982
  by (simp add: ac_simps) (simp only: flip: take_bit_take_bit)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   983
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   984
context
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   985
  includes bit_operations_syntax
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   986
begin
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   987
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   988
instance proof
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   989
  fix v w :: \<open>'a word\<close> and n m :: nat
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79031
diff changeset
   990
  show \<open>NOT v = - v - 1\<close>
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79031
diff changeset
   991
    by transfer (simp add: not_eq_complement)
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   992
  show \<open>v AND w = of_bool (odd v \<and> odd w) + 2 * (v div 2 AND w div 2)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   993
    apply transfer
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   994
    by (rule bit_eqI) (auto simp: even_bit_succ_iff bit_simps bit_0 simp flip: bit_Suc)
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   995
  show \<open>v OR w = of_bool (odd v \<or> odd w) + 2 * (v div 2 OR w div 2)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   996
    apply transfer
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
   997
    by (rule bit_eqI) (auto simp: even_bit_succ_iff bit_simps bit_0 simp flip: bit_Suc)
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   998
  show \<open>v XOR w = of_bool (odd v \<noteq> odd w) + 2 * (v div 2 XOR w div 2)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   999
    apply transfer
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1000
    apply (rule bit_eqI)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1001
    subgoal for k l n
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1002
      apply (cases n)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1003
       apply (auto simp: even_bit_succ_iff bit_simps bit_0 even_xor_iff simp flip: bit_Suc)
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1004
      done
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1005
    done
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1006
  show \<open>mask n = 2 ^ n - (1 :: 'a word)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1007
    by transfer (simp flip: mask_eq_exp_minus_1)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1008
  show \<open>set_bit n v = v OR push_bit n 1\<close>
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79481
diff changeset
  1009
    by transfer (simp add: set_bit_eq_or)
1e19abf373ac streamlined type class specification
haftmann
parents: 79481
diff changeset
  1010
  show \<open>unset_bit n v = (v OR push_bit n 1) XOR push_bit n 1\<close>
1e19abf373ac streamlined type class specification
haftmann
parents: 79481
diff changeset
  1011
    by transfer (simp add: unset_bit_eq_or_xor)
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1012
  show \<open>flip_bit n v = v XOR push_bit n 1\<close>
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79481
diff changeset
  1013
    by transfer (simp add: flip_bit_eq_xor)
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1014
  show \<open>push_bit n v = v * 2 ^ n\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1015
    by transfer (simp add: push_bit_eq_mult)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1016
  show \<open>drop_bit n v = v div 2 ^ n\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1017
    by transfer (simp add: drop_bit_take_bit flip: drop_bit_eq_div)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1018
  show \<open>take_bit n v = v mod 2 ^ n\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1019
    by transfer (simp flip: take_bit_eq_mod)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1020
qed
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1021
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  1022
end
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1023
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1024
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1025
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1026
lemma [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1027
  \<open>push_bit n w = w * 2 ^ n\<close> for w :: \<open>'a::len word\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1028
  by (fact push_bit_eq_mult)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1029
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1030
lemma [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1031
  \<open>Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1032
  by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1033
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1034
lemma [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1035
  \<open>Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1036
  for w :: \<open>'a::len word\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1037
  by transfer (simp add: not_le not_less ac_simps min_absorb2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1038
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1039
lemma [code_abbrev]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1040
  \<open>push_bit n 1 = (2 :: 'a::len word) ^ n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1041
  by (fact push_bit_of_1)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1042
74391
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1043
context
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1044
  includes bit_operations_syntax
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1045
begin
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1046
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1047
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1048
  \<open>NOT w = Word.of_int (NOT (Word.the_int w))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1049
  for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1050
  by transfer (simp add: take_bit_not_take_bit) 
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1051
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1052
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1053
  \<open>Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\<close>
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1054
  by transfer simp
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1055
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1056
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1057
  \<open>Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1058
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1059
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1060
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1061
  \<open>Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1062
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1063
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1064
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1065
  \<open>Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1066
  by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1067
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1068
lemma [code]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1069
  \<open>set_bit n w = w OR push_bit n 1\<close> for w :: \<open>'a::len word\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1070
  by (fact set_bit_eq_or)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1071
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1072
lemma [code]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1073
  \<open>unset_bit n w = w AND NOT (push_bit n 1)\<close> for w :: \<open>'a::len word\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1074
  by (fact unset_bit_eq_and_not)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1075
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1076
lemma [code]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1077
  \<open>flip_bit n w = w XOR push_bit n 1\<close> for w :: \<open>'a::len word\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1078
  by (fact flip_bit_eq_xor)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1079
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1080
context
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1081
  includes lifting_syntax
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1082
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1083
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1084
lemma set_bit_word_transfer [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1085
  \<open>((=) ===> pcr_word ===> pcr_word) set_bit set_bit\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1086
  by (unfold set_bit_def) transfer_prover
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1087
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1088
lemma unset_bit_word_transfer [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1089
  \<open>((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1090
  by (unfold unset_bit_def) transfer_prover
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1091
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1092
lemma flip_bit_word_transfer [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1093
  \<open>((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1094
  by (unfold flip_bit_def) transfer_prover
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1095
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1096
lemma signed_take_bit_word_transfer [transfer_rule]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1097
  \<open>((=) ===> pcr_word ===> pcr_word)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1098
    (\<lambda>n k. signed_take_bit n (take_bit LENGTH('a::len) k))
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1099
    (signed_take_bit :: nat \<Rightarrow> 'a word \<Rightarrow> 'a word)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1100
proof -
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1101
  let ?K = \<open>\<lambda>n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \<and> bit k n) * NOT (mask n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1102
  let ?W = \<open>\<lambda>n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1103
  have \<open>((=) ===> pcr_word ===> pcr_word) ?K ?W\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1104
    by transfer_prover
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1105
  also have \<open>?K = (\<lambda>n k. signed_take_bit n (take_bit LENGTH('a::len) k))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1106
    by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1107
  also have \<open>?W = signed_take_bit\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1108
    by (simp add: fun_eq_iff signed_take_bit_def)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1109
  finally show ?thesis .
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1110
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1111
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1112
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1113
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1114
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1115
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1116
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1117
subsection \<open>Conversions including casts\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1118
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1119
subsubsection \<open>Generic unsigned conversion\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1120
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1121
context semiring_bits
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1122
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1123
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  1124
lemma bit_unsigned_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1125
  \<open>bit (unsigned w) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit w n\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1126
  for w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1127
  by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1128
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1129
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1130
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1131
lemma possible_bit_word[simp]:
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1132
  \<open>possible_bit TYPE(('a :: len) word) m \<longleftrightarrow> m < LENGTH('a)\<close>
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1133
  by (simp add: possible_bit_def linorder_not_le)
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1134
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1135
context semiring_bit_operations
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1136
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1137
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1138
lemma unsigned_minus_1_eq_mask:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1139
  \<open>unsigned (- 1 :: 'b::len word) = mask LENGTH('b)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1140
  by (transfer fixing: mask) (simp add: nat_mask_eq of_nat_mask_eq)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1141
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1142
lemma unsigned_push_bit_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1143
  \<open>unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1144
  for w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1145
proof (rule bit_eqI)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1146
  fix m
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1147
  assume \<open>possible_bit TYPE('a) m\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1148
  show \<open>bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1149
  proof (cases \<open>n \<le> m\<close>)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1150
    case True
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1151
    with \<open>possible_bit TYPE('a) m\<close> have \<open>possible_bit TYPE('a) (m - n)\<close>
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1152
      by (simp add: possible_bit_less_imp)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1153
    with True show ?thesis
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1154
      by (simp add: bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff not_le ac_simps)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1155
  next
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1156
    case False
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1157
    then show ?thesis
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1158
      by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1159
  qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1160
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1161
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1162
lemma unsigned_take_bit_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1163
  \<open>unsigned (take_bit n w) = take_bit n (unsigned w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1164
  for w :: \<open>'b::len word\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1165
  by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Bit_Operations.bit_take_bit_iff)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1166
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1167
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1168
78955
74147aa81dbb more specific name for type class
haftmann
parents: 77812
diff changeset
  1169
context linordered_euclidean_semiring_bit_operations
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1170
begin
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1171
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1172
lemma unsigned_drop_bit_eq:
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1173
  \<open>unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\<close>
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1174
  for w :: \<open>'b::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1175
  by (rule bit_eqI) (auto simp: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Bit_Operations.bit_drop_bit_eq possible_bit_def dest: bit_imp_le_length)
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1176
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1177
end
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1178
73853
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  1179
lemma ucast_drop_bit_eq:
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  1180
  \<open>ucast (drop_bit n w) = drop_bit n (ucast w :: 'b::len word)\<close>
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  1181
  if \<open>LENGTH('a) \<le> LENGTH('b)\<close> for w :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1182
  by (rule bit_word_eqI) (use that in \<open>auto simp: bit_unsigned_iff bit_drop_bit_eq dest: bit_imp_le_length\<close>)
73853
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  1183
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1184
context semiring_bit_operations
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1185
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1186
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1187
context
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1188
  includes bit_operations_syntax
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1189
begin
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1190
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1191
lemma unsigned_and_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1192
  \<open>unsigned (v AND w) = unsigned v AND unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1193
  for v w :: \<open>'b::len word\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1194
  by (simp add: bit_eq_iff bit_simps)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1195
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1196
lemma unsigned_or_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1197
  \<open>unsigned (v OR w) = unsigned v OR unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1198
  for v w :: \<open>'b::len word\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1199
  by (simp add: bit_eq_iff bit_simps)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1200
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1201
lemma unsigned_xor_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1202
  \<open>unsigned (v XOR w) = unsigned v XOR unsigned w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1203
  for v w :: \<open>'b::len word\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1204
  by (simp add: bit_eq_iff bit_simps)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1205
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1206
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1207
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1208
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1209
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1210
context ring_bit_operations
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1211
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1212
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1213
context
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1214
  includes bit_operations_syntax
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1215
begin
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1216
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1217
lemma unsigned_not_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1218
  \<open>unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1219
  for w :: \<open>'b::len word\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1220
  by (simp add: bit_eq_iff bit_simps)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1221
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1222
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1223
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1224
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1225
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  1226
context linordered_euclidean_semiring
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1227
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1228
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  1229
lemma unsigned_greater_eq [simp]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1230
  \<open>0 \<le> unsigned w\<close> for w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1231
  by (transfer fixing: less_eq) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1232
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  1233
lemma unsigned_less [simp]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1234
  \<open>unsigned w < 2 ^ LENGTH('b)\<close> for w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1235
  by (transfer fixing: less) simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1236
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1237
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1238
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1239
context linordered_semidom
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1240
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1241
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1242
lemma word_less_eq_iff_unsigned:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1243
  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1244
  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1245
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1246
lemma word_less_iff_unsigned:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1247
  "a < b \<longleftrightarrow> unsigned a < unsigned b"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1248
  by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative])
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1249
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1250
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1251
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1252
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1253
subsubsection \<open>Generic signed conversion\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1254
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1255
context ring_bit_operations
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1256
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1257
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  1258
lemma bit_signed_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1259
  \<open>bit (signed w) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit w (min (LENGTH('b) - Suc 0) n)\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1260
  for w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1261
  by (transfer fixing: bit)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1262
    (auto simp: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1263
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1264
lemma signed_push_bit_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1265
  \<open>signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1266
  for w :: \<open>'b::len word\<close>
82524
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1267
  by (simp add: bit_eq_iff bit_simps possible_bit_less_imp min_less_iff_disj) auto
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1268
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1269
lemma signed_take_bit_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1270
  \<open>signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1271
  for w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1272
  by (transfer fixing: take_bit; cases \<open>LENGTH('b)\<close>)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1273
    (auto simp: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1274
74391
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1275
context
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1276
  includes bit_operations_syntax
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1277
begin
930047942f46 repaired slip
haftmann
parents: 74309
diff changeset
  1278
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1279
lemma signed_not_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1280
  \<open>signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1281
  for w :: \<open>'b::len word\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1282
  by (simp add: bit_eq_iff bit_simps possible_bit_less_imp min_less_iff_disj)
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1283
    (auto simp: min_def)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1284
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1285
lemma signed_and_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1286
  \<open>signed (v AND w) = signed v AND signed w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1287
  for v w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1288
  by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1289
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1290
lemma signed_or_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1291
  \<open>signed (v OR w) = signed v OR signed w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1292
  for v w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1293
  by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1294
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1295
lemma signed_xor_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1296
  \<open>signed (v XOR w) = signed v XOR signed w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1297
  for v w :: \<open>'b::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1298
  by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1299
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1300
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1301
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1302
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1303
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1304
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1305
subsubsection \<open>More\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1306
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1307
lemma sint_greater_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1308
  \<open>- (2 ^ (LENGTH('a) - Suc 0)) \<le> sint w\<close> for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1309
proof (cases \<open>bit w (LENGTH('a) - Suc 0)\<close>)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1310
  case True
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1311
  then show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1312
    by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1313
next
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1314
  have *: \<open>- (2 ^ (LENGTH('a) - Suc 0)) \<le> (0::int)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1315
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1316
  case False
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1317
  then show ?thesis
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1318
    by transfer (auto simp: signed_take_bit_eq intro: order_trans *)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1319
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1320
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1321
lemma sint_less:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1322
  \<open>sint w < 2 ^ (LENGTH('a) - Suc 0)\<close> for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1323
  by (cases \<open>bit w (LENGTH('a) - Suc 0)\<close>; transfer)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1324
    (simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1325
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1326
lemma uint_div_distrib:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1327
  \<open>uint (v div w) = uint v div uint w\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1328
  by (metis int_ops(8) of_nat_unat unat_div_distrib)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1329
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1330
lemma unat_drop_bit_eq:
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1331
  \<open>unat (drop_bit n w) = drop_bit n (unat w)\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1332
  by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq)
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1333
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1334
lemma uint_mod_distrib:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1335
  \<open>uint (v mod w) = uint v mod uint w\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1336
  by (metis int_ops(9) of_nat_unat unat_mod_distrib)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1337
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1338
context semiring_bit_operations
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1339
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1340
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1341
lemma unsigned_ucast_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1342
  \<open>unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1343
  for w :: \<open>'b::len word\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1344
  by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff not_le)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1345
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1346
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1347
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1348
context ring_bit_operations
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1349
begin
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1350
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1351
lemma signed_ucast_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1352
  \<open>signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1353
  for w :: \<open>'b::len word\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1354
  by (simp add: bit_eq_iff bit_simps min_less_iff_disj)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1355
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1356
lemma signed_scast_eq:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1357
  \<open>signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1358
  for w :: \<open>'b::len word\<close>
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
  1359
  by (simp add: bit_eq_iff bit_simps min_less_iff_disj)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1360
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1361
end
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1362
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1363
lemma uint_nonnegative: "0 \<le> uint w"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1364
  by (fact unsigned_greater_eq)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1365
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1366
lemma uint_bounded: "uint w < 2 ^ LENGTH('a)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1367
  for w :: "'a::len word"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1368
  by (fact unsigned_less)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1369
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1370
lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1371
  for w :: "'a::len word"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1372
  by transfer (simp add: take_bit_eq_mod)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1373
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1374
lemma word_uint_eqI: "uint a = uint b \<Longrightarrow> a = b"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1375
  by (fact unsigned_word_eqI)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1376
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1377
lemma word_uint_eq_iff: "a = b \<longleftrightarrow> uint a = uint b"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1378
  by (fact word_eq_iff_unsigned)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1379
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1380
lemma uint_word_of_int_eq:
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1381
  \<open>uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1382
  by transfer rule
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1383
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1384
lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1385
  by (simp add: uint_word_of_int_eq take_bit_eq_mod)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1386
  
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1387
lemma word_of_int_uint: "word_of_int (uint w) = w"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1388
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1389
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1390
lemma word_div_def [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1391
  "a div b = word_of_int (uint a div uint b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1392
  by transfer rule
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1393
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1394
lemma word_mod_def [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1395
  "a mod b = word_of_int (uint a mod uint b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1396
  by transfer rule
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1397
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1398
lemma split_word_all: "(\<And>x::'a::len word. PROP P x) \<equiv> (\<And>x. PROP P (word_of_int x))"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1399
proof
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1400
  fix x :: "'a word"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1401
  assume "\<And>x. PROP P (word_of_int x)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1402
  then have "PROP P (word_of_int (uint x))" .
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1403
  then show "PROP P x"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1404
    by (simp only: word_of_int_uint)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1405
qed
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1406
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1407
lemma sint_uint:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1408
  \<open>sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\<close>
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1409
  for w :: \<open>'a::len word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1410
  by (cases \<open>LENGTH('a)\<close>; transfer) (simp_all add: signed_take_bit_take_bit)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1411
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1412
lemma unat_eq_nat_uint:
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1413
  \<open>unat w = nat (uint w)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1414
  by simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1415
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1416
lemma ucast_eq:
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1417
  \<open>ucast w = word_of_int (uint w)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1418
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1419
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1420
lemma scast_eq:
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1421
  \<open>scast w = word_of_int (sint w)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1422
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1423
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1424
lemma uint_0_eq:
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1425
  \<open>uint 0 = 0\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1426
  by (fact unsigned_0)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1427
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1428
lemma uint_1_eq:
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1429
  \<open>uint 1 = 1\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1430
  by (fact unsigned_1)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1431
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1432
lemma word_m1_wi: "- 1 = word_of_int (- 1)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1433
  by simp
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1434
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1435
lemma uint_0_iff: "uint x = 0 \<longleftrightarrow> x = 0"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1436
  by (auto simp: unsigned_word_eqI)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1437
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1438
lemma unat_0_iff: "unat x = 0 \<longleftrightarrow> x = 0"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1439
  by (auto simp: unsigned_word_eqI)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1440
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1441
lemma unat_0: "unat 0 = 0"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1442
  by (fact unsigned_0)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1443
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1444
lemma unat_gt_0: "0 < unat x \<longleftrightarrow> x \<noteq> 0"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1445
  by (auto simp: unat_0_iff [symmetric])
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1446
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1447
lemma ucast_0: "ucast 0 = 0"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1448
  by (fact unsigned_0)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1449
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1450
lemma sint_0: "sint 0 = 0"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1451
  by (fact signed_0)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1452
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1453
lemma scast_0: "scast 0 = 0"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1454
  by (fact signed_0)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1455
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1456
lemma sint_n1: "sint (- 1) = - 1"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1457
  by (fact signed_minus_1)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1458
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1459
lemma scast_n1: "scast (- 1) = - 1"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1460
  by (fact signed_minus_1)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1461
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1462
lemma uint_1: "uint (1::'a::len word) = 1"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1463
  by (fact uint_1_eq)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1464
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1465
lemma unat_1: "unat (1::'a::len word) = 1"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1466
  by (fact unsigned_1)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1467
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1468
lemma ucast_1: "ucast (1::'a::len word) = 1"
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1469
  by (fact unsigned_1)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1470
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1471
instantiation word :: (len) size
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1472
begin
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1473
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1474
lift_definition size_word :: \<open>'a word \<Rightarrow> nat\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1475
  is \<open>\<lambda>_. LENGTH('a)\<close> ..
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1476
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1477
instance ..
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1478
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1479
end
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1480
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1481
lemma word_size [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1482
  \<open>size w = LENGTH('a)\<close> for w :: \<open>'a::len word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1483
  by (fact size_word.rep_eq)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1484
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1485
lemma word_size_gt_0 [iff]: "0 < size w"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1486
  for w :: "'a::len word"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1487
  by (simp add: word_size)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1488
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1489
lemmas lens_gt_0 = word_size_gt_0 len_gt_0
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1490
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1491
lemma lens_not_0 [iff]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1492
  \<open>size w \<noteq> 0\<close> for  w :: \<open>'a::len word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1493
  by auto
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1494
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1495
lift_definition source_size :: \<open>('a::len word \<Rightarrow> 'b) \<Rightarrow> nat\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1496
  is \<open>\<lambda>_. LENGTH('a)\<close> .
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1497
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1498
lift_definition target_size :: \<open>('a \<Rightarrow> 'b::len word) \<Rightarrow> nat\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1499
  is \<open>\<lambda>_. LENGTH('b)\<close> ..
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1500
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1501
lift_definition is_up :: \<open>('a::len word \<Rightarrow> 'b::len word) \<Rightarrow> bool\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1502
  is \<open>\<lambda>_. LENGTH('a) \<le> LENGTH('b)\<close> ..
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1503
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1504
lift_definition is_down :: \<open>('a::len word \<Rightarrow> 'b::len word) \<Rightarrow> bool\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1505
  is \<open>\<lambda>_. LENGTH('a) \<ge> LENGTH('b)\<close> ..
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1506
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1507
lemma is_up_eq:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1508
  \<open>is_up f \<longleftrightarrow> source_size f \<le> target_size f\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1509
  for f :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1510
  by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1511
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1512
lemma is_down_eq:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1513
  \<open>is_down f \<longleftrightarrow> target_size f \<le> source_size f\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1514
  for f :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1515
  by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1516
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1517
lift_definition word_int_case :: \<open>(int \<Rightarrow> 'b) \<Rightarrow> 'a::len word \<Rightarrow> 'b\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1518
  is \<open>\<lambda>f. f \<circ> take_bit LENGTH('a)\<close> by simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1519
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1520
lemma word_int_case_eq_uint [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1521
  \<open>word_int_case f w = f (uint w)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1522
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1523
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1524
translations
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1525
  "case x of XCONST of_int y \<Rightarrow> b" \<rightleftharpoons> "CONST word_int_case (\<lambda>y. b) x"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1526
  "case x of (XCONST of_int :: 'a) y \<Rightarrow> b" \<rightharpoonup> "CONST word_int_case (\<lambda>y. b) x"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1527
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1528
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1529
subsection \<open>Arithmetic operations\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1530
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1531
lemma div_word_self:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1532
  \<open>w div w = 1\<close> if \<open>w \<noteq> 0\<close> for w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1533
  using that by transfer simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1534
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1535
lemma mod_word_self [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1536
  \<open>w mod w = 0\<close> for w :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1537
  by (simp add: word_mod_def)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1538
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1539
lemma div_word_less:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1540
  \<open>w div v = 0\<close> if \<open>w < v\<close> for w v :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1541
  using that by transfer simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1542
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1543
lemma mod_word_less:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1544
  \<open>w mod v = w\<close> if \<open>w < v\<close> for w v :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1545
  using div_mult_mod_eq [of w v] using that by (simp add: div_word_less)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1546
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1547
lemma div_word_one [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1548
  \<open>1 div w = of_bool (w = 1)\<close> for w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1549
proof transfer
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1550
  fix k :: int
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1551
  show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) 1 div take_bit LENGTH('a) k) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1552
         take_bit LENGTH('a) (of_bool (take_bit LENGTH('a) k = take_bit LENGTH('a) 1))\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1553
    using take_bit_nonnegative [of \<open>LENGTH('a)\<close> k]
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1554
    by (smt (verit, best) div_by_1 of_bool_eq take_bit_of_0 take_bit_of_1 zdiv_eq_0_iff)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1555
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1556
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1557
lemma mod_word_one [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1558
  \<open>1 mod w = 1 - w * of_bool (w = 1)\<close> for w :: \<open>'a::len word\<close>
75087
f3fcc7c5a0db Avoid overaggresive splitting.
haftmann
parents: 75085
diff changeset
  1559
  using div_mult_mod_eq [of 1 w] by auto
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1560
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1561
lemma div_word_by_minus_1_eq [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1562
  \<open>w div - 1 = of_bool (w = - 1)\<close> for w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1563
  by (auto intro: div_word_less simp add: div_word_self word_order.not_eq_extremum)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1564
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1565
lemma mod_word_by_minus_1_eq [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1566
  \<open>w mod - 1 = w * of_bool (w < - 1)\<close> for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1567
  using mod_word_less word_order.not_eq_extremum by fastforce
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1568
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1569
text \<open>Legacy theorems:\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1570
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1571
lemma word_add_def [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1572
  "a + b = word_of_int (uint a + uint b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1573
  by transfer (simp add: take_bit_add)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1574
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1575
lemma word_sub_wi [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1576
  "a - b = word_of_int (uint a - uint b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1577
  by transfer (simp add: take_bit_diff)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1578
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1579
lemma word_mult_def [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1580
  "a * b = word_of_int (uint a * uint b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1581
  by transfer (simp add: take_bit_eq_mod mod_simps)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1582
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1583
lemma word_minus_def [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1584
  "- a = word_of_int (- uint a)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1585
  by transfer (simp add: take_bit_minus)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1586
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1587
lemma word_0_wi:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1588
  "0 = word_of_int 0"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1589
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1590
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1591
lemma word_1_wi:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1592
  "1 = word_of_int 1"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1593
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1594
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1595
lift_definition word_succ :: "'a::len word \<Rightarrow> 'a word" is "\<lambda>x. x + 1"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1596
  by (auto simp: take_bit_eq_mod intro: mod_add_cong)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1597
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1598
lift_definition word_pred :: "'a::len word \<Rightarrow> 'a word" is "\<lambda>x. x - 1"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1599
  by (auto simp: take_bit_eq_mod intro: mod_diff_cong)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1600
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1601
lemma word_succ_alt [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1602
  "word_succ a = word_of_int (uint a + 1)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1603
  by transfer (simp add: take_bit_eq_mod mod_simps)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1604
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1605
lemma word_pred_alt [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1606
  "word_pred a = word_of_int (uint a - 1)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1607
  by transfer (simp add: take_bit_eq_mod mod_simps)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1608
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1609
lemmas word_arith_wis = 
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1610
  word_add_def word_sub_wi word_mult_def
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1611
  word_minus_def word_succ_alt word_pred_alt
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1612
  word_0_wi word_1_wi
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1613
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1614
lemma wi_homs:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1615
  shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1616
    and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1617
    and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1618
    and wi_hom_neg: "- word_of_int a = word_of_int (- a)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1619
    and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1620
    and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)"
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1621
  by (transfer, simp)+
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1622
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1623
lemmas wi_hom_syms = wi_homs [symmetric]
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1624
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1625
lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1626
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1627
lemmas word_of_int_hom_syms = word_of_int_homs [symmetric]
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1628
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1629
lemma double_eq_zero_iff:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1630
  \<open>2 * a = 0 \<longleftrightarrow> a = 0 \<or> a = 2 ^ (LENGTH('a) - Suc 0)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1631
  for a :: \<open>'a::len word\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1632
proof -
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1633
  define n where \<open>n = LENGTH('a) - Suc 0\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1634
  then have *: \<open>LENGTH('a) = Suc n\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1635
    by simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1636
  have \<open>a = 0\<close> if \<open>2 * a = 0\<close> and \<open>a \<noteq> 2 ^ (LENGTH('a) - Suc 0)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1637
    using that by transfer
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1638
      (auto simp: take_bit_eq_0_iff take_bit_eq_mod *)
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1639
  moreover have \<open>2 ^ LENGTH('a) = (0 :: 'a word)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1640
    by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1641
  then have \<open>2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1642
    by (simp add: *)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1643
  ultimately show ?thesis
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1644
    by auto
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1645
qed
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1646
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1647
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1648
subsection \<open>Ordering\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1649
82376
haftmann
parents: 82367
diff changeset
  1650
instance word :: (len) wellorder
haftmann
parents: 82367
diff changeset
  1651
proof
haftmann
parents: 82367
diff changeset
  1652
  fix P :: "'a word \<Rightarrow> bool" and a
haftmann
parents: 82367
diff changeset
  1653
  assume *: "(\<And>b. (\<And>a. a < b \<Longrightarrow> P a) \<Longrightarrow> P b)"
haftmann
parents: 82367
diff changeset
  1654
  have "wf (measure unat)" ..
haftmann
parents: 82367
diff changeset
  1655
  moreover have "{(a, b :: ('a::len) word). a < b} \<subseteq> measure unat"
haftmann
parents: 82367
diff changeset
  1656
    by (auto simp add: word_less_iff_unsigned [where ?'a = nat])
haftmann
parents: 82367
diff changeset
  1657
  ultimately have "wf {(a, b :: ('a::len) word). a < b}"
haftmann
parents: 82367
diff changeset
  1658
    by (rule wf_subset)
haftmann
parents: 82367
diff changeset
  1659
  then show "P a" using *
haftmann
parents: 82367
diff changeset
  1660
    by induction blast
haftmann
parents: 82367
diff changeset
  1661
qed
haftmann
parents: 82367
diff changeset
  1662
haftmann
parents: 82367
diff changeset
  1663
lemma word_m1_ge [simp]: (* FIXME: delete *)
haftmann
parents: 82367
diff changeset
  1664
  "word_pred 0 \<ge> y"
haftmann
parents: 82367
diff changeset
  1665
  by transfer (simp add: mask_eq_exp_minus_1)
haftmann
parents: 82367
diff changeset
  1666
haftmann
parents: 82367
diff changeset
  1667
lemma word_less_alt:
haftmann
parents: 82367
diff changeset
  1668
  "a < b \<longleftrightarrow> uint a < uint b"
haftmann
parents: 82367
diff changeset
  1669
  by (fact word_less_def)
haftmann
parents: 82367
diff changeset
  1670
haftmann
parents: 82367
diff changeset
  1671
lemma word_zero_le [simp]:
haftmann
parents: 82367
diff changeset
  1672
  "0 \<le> y" for y :: "'a::len word"
haftmann
parents: 82367
diff changeset
  1673
  by (fact word_coorder.extremum)
haftmann
parents: 82367
diff changeset
  1674
haftmann
parents: 82367
diff changeset
  1675
lemma word_n1_ge [simp]:
haftmann
parents: 82367
diff changeset
  1676
  "y \<le> - 1" for y :: "'a::len word"
haftmann
parents: 82367
diff changeset
  1677
  by (fact word_order.extremum)
haftmann
parents: 82367
diff changeset
  1678
haftmann
parents: 82367
diff changeset
  1679
lemmas word_not_simps [simp] =
haftmann
parents: 82367
diff changeset
  1680
  word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD]
haftmann
parents: 82367
diff changeset
  1681
haftmann
parents: 82367
diff changeset
  1682
lemma word_gt_0:
haftmann
parents: 82367
diff changeset
  1683
  "0 < y \<longleftrightarrow> 0 \<noteq> y"
haftmann
parents: 82367
diff changeset
  1684
  for y :: "'a::len word"
haftmann
parents: 82367
diff changeset
  1685
  by (simp add: less_le)
haftmann
parents: 82367
diff changeset
  1686
haftmann
parents: 82367
diff changeset
  1687
lemma word_gt_0_no [simp]:
haftmann
parents: 82367
diff changeset
  1688
  \<open>(0 :: 'a::len word) < numeral y \<longleftrightarrow> (0 :: 'a::len word) \<noteq> numeral y\<close>
haftmann
parents: 82367
diff changeset
  1689
  by (fact word_gt_0)
haftmann
parents: 82367
diff changeset
  1690
haftmann
parents: 82367
diff changeset
  1691
lemma word_le_nat_alt:
haftmann
parents: 82367
diff changeset
  1692
  "a \<le> b \<longleftrightarrow> unat a \<le> unat b"
haftmann
parents: 82367
diff changeset
  1693
  by transfer (simp add: nat_le_eq_zle)
haftmann
parents: 82367
diff changeset
  1694
haftmann
parents: 82367
diff changeset
  1695
lemma word_less_nat_alt:
haftmann
parents: 82367
diff changeset
  1696
  "a < b \<longleftrightarrow> unat a < unat b"
haftmann
parents: 82367
diff changeset
  1697
  by transfer (auto simp: less_le [of 0])
haftmann
parents: 82367
diff changeset
  1698
haftmann
parents: 82367
diff changeset
  1699
lemmas unat_mono = word_less_nat_alt [THEN iffD1]
haftmann
parents: 82367
diff changeset
  1700
haftmann
parents: 82367
diff changeset
  1701
lemma wi_less:
haftmann
parents: 82367
diff changeset
  1702
  "(word_of_int n < (word_of_int m :: 'a::len word)) =
haftmann
parents: 82367
diff changeset
  1703
    (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))"
haftmann
parents: 82367
diff changeset
  1704
  by (simp add: uint_word_of_int word_less_def)
haftmann
parents: 82367
diff changeset
  1705
haftmann
parents: 82367
diff changeset
  1706
lemma wi_le:
haftmann
parents: 82367
diff changeset
  1707
  "(word_of_int n \<le> (word_of_int m :: 'a::len word)) =
haftmann
parents: 82367
diff changeset
  1708
    (n mod 2 ^ LENGTH('a) \<le> m mod 2 ^ LENGTH('a))"
haftmann
parents: 82367
diff changeset
  1709
  by (simp add: uint_word_of_int word_le_def)
haftmann
parents: 82367
diff changeset
  1710
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1711
lift_definition word_sle :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> bool\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1712
  is \<open>\<lambda>k l. signed_take_bit (LENGTH('a) - Suc 0) k \<le> signed_take_bit (LENGTH('a) - Suc 0) l\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1713
  by (simp flip: signed_take_bit_decr_length_iff)
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1714
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1715
lift_definition word_sless :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> bool\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1716
  is \<open>\<lambda>k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\<close>
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1717
  by (simp flip: signed_take_bit_decr_length_iff)
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1718
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1719
notation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80777
diff changeset
  1720
  word_sle    (\<open>'(\<le>s')\<close>) and
81142
6ad2c917dd2e more inner-syntax markup;
wenzelm
parents: 80914
diff changeset
  1721
  word_sle    (\<open>(\<open>notation=\<open>infix \<le>s\<close>\<close>_/ \<le>s _)\<close>  [51, 51] 50) and
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80777
diff changeset
  1722
  word_sless  (\<open>'(<s')\<close>) and
81142
6ad2c917dd2e more inner-syntax markup;
wenzelm
parents: 80914
diff changeset
  1723
  word_sless  (\<open>(\<open>notation=\<open>infix <s\<close>\<close>_/ <s _)\<close>  [51, 51] 50)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1724
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1725
notation (input)
81142
6ad2c917dd2e more inner-syntax markup;
wenzelm
parents: 80914
diff changeset
  1726
  word_sle    (\<open>(\<open>notation=\<open>infix <=s\<close>\<close>_/ <=s _)\<close>  [51, 51] 50)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1727
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1728
lemma word_sle_eq [code]:
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1729
  \<open>a <=s b \<longleftrightarrow> sint a \<le> sint b\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1730
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1731
82376
haftmann
parents: 82367
diff changeset
  1732
lemma word_sless_alt [code]:
haftmann
parents: 82367
diff changeset
  1733
  "a <s b \<longleftrightarrow> sint a < sint b"
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1734
  by transfer simp
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1735
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1736
lemma signed_ordering: \<open>ordering word_sle word_sless\<close>
82376
haftmann
parents: 82367
diff changeset
  1737
  by (standard; transfer) (auto simp flip: signed_take_bit_decr_length_iff)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1738
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1739
lemma signed_linorder: \<open>class.linorder word_sle word_sless\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1740
  by (standard; transfer) (auto simp: signed_take_bit_decr_length_iff)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1741
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1742
interpretation signed: linorder word_sle word_sless
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1743
  by (fact signed_linorder)
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1744
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1745
lemma word_sless_eq:
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1746
  \<open>x <s y \<longleftrightarrow> x <=s y \<and> x \<noteq> y\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1747
  by (fact signed.less_le)
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1748
82524
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1749
lemma minus_1_sless_0 [simp]:
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1750
  \<open>- 1 <s 0\<close>
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1751
  by transfer simp
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1752
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1753
lemma not_0_sless_minus_1 [simp]:
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1754
  \<open>\<not> 0 <s - 1\<close>
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1755
  by transfer simp
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1756
82524
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1757
lemma minus_1_sless_eq_0 [simp]:
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1758
  \<open>- 1 \<le>s 0\<close>
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1759
  by transfer simp
df5b2785abd6 more lemmas
haftmann
parents: 82523
diff changeset
  1760
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1761
lemma not_0_sless_eq_minus_1 [simp]:
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1762
  \<open>\<not> 0 \<le>s - 1\<close>
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1763
  by transfer simp
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1764
72244
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1765
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1766
subsection \<open>Bit-wise operations\<close>
4b011fa5e83b restructured
haftmann
parents: 72243
diff changeset
  1767
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1768
context
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1769
  includes bit_operations_syntax
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1770
begin
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  1771
72227
0f3d24dc197f more on conversions
haftmann
parents: 72130
diff changeset
  1772
lemma take_bit_word_eq_self:
0f3d24dc197f more on conversions
haftmann
parents: 72130
diff changeset
  1773
  \<open>take_bit n w = w\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72130
diff changeset
  1774
  using that by transfer simp
0f3d24dc197f more on conversions
haftmann
parents: 72130
diff changeset
  1775
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  1776
lemma take_bit_length_eq [simp]:
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  1777
  \<open>take_bit LENGTH('a) w = w\<close> for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1778
  by (simp add: nat_le_linear take_bit_word_eq_self)
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  1779
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1780
lemma bit_word_of_int_iff:
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1781
  \<open>bit (word_of_int k :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> bit k n\<close>
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1782
  by (simp add: bit_simps)
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1783
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1784
lemma bit_uint_iff:
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1785
  \<open>bit (uint w) n \<longleftrightarrow> n < LENGTH('a) \<and> bit w n\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1786
    for w :: \<open>'a::len word\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1787
  by transfer (simp add: bit_take_bit_iff)
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1788
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1789
lemma bit_sint_iff:
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1790
  \<open>bit (sint w) n \<longleftrightarrow> n \<ge> LENGTH('a) \<and> bit w (LENGTH('a) - 1) \<or> bit w n\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1791
  for w :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1792
  by transfer (auto simp: bit_signed_take_bit_iff min_def le_less not_less)
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1793
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1794
lemma bit_word_ucast_iff:
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1795
  \<open>bit (ucast w :: 'b::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> n < LENGTH('b) \<and> bit w n\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1796
  for w :: \<open>'a::len word\<close>
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1797
  by (auto simp add: bit_unsigned_iff bit_imp_le_length)
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1798
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1799
lemma bit_word_scast_iff:
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1800
  \<open>bit (scast w :: 'b::len word) n \<longleftrightarrow>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1801
    n < LENGTH('b) \<and> (bit w n \<or> LENGTH('a) \<le> n \<and> bit w (LENGTH('a) - Suc 0))\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  1802
  for w :: \<open>'a::len word\<close>
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1803
  by (auto simp add: bit_signed_iff bit_imp_le_length min_def le_less dest: bit_imp_possible_bit)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1804
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  1805
lemma bit_word_iff_drop_bit_and [code]:
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  1806
  \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close> for a :: \<open>'a::len word\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  1807
  by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1808
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1809
lemma
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1810
  word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))"
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  1811
    and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  1812
    and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  1813
    and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)"
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1814
  by (transfer, simp add: take_bit_not_take_bit)+
47374
9475d524bafb set up and use lift_definition for word operations
huffman
parents: 47372
diff changeset
  1815
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1816
definition even_word :: \<open>'a::len word \<Rightarrow> bool\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1817
  where [code_abbrev]: \<open>even_word = even\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1818
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1819
lemma even_word_iff [code]:
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1820
  \<open>even_word a \<longleftrightarrow> a AND 1 = 0\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1821
  by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def)
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  1822
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1823
lemma map_bit_range_eq_if_take_bit_eq:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1824
  \<open>map (bit k) [0..<n] = map (bit l) [0..<n]\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1825
  if \<open>take_bit n k = take_bit n l\<close> for k l :: int
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1826
  using that 
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1827
proof (induction n arbitrary: k l)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1828
  case 0
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1829
  then show ?case
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1830
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1831
next
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1832
  case (Suc n)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1833
  from Suc.prems have \<open>take_bit n (k div 2) = take_bit n (l div 2)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1834
    by (simp add: take_bit_Suc)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1835
  then have \<open>map (bit (k div 2)) [0..<n] = map (bit (l div 2)) [0..<n]\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1836
    by (rule Suc.IH)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1837
  moreover have \<open>bit (r div 2) = bit r \<circ> Suc\<close> for r :: int
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1838
    by (simp add: fun_eq_iff bit_Suc)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1839
  moreover from Suc.prems have \<open>even k \<longleftrightarrow> even l\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1840
    by (metis Zero_neq_Suc even_take_bit_eq)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1841
  ultimately show ?case
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74592
diff changeset
  1842
    by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) (simp add: bit_0)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1843
qed
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1844
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1845
lemma
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1846
  take_bit_word_Bit0_eq [simp]: \<open>take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1847
    = 2 * take_bit (pred_numeral n) (numeral m)\<close> (is ?P)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1848
  and take_bit_word_Bit1_eq [simp]: \<open>take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1849
    = 1 + 2 * take_bit (pred_numeral n) (numeral m)\<close> (is ?Q)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1850
  and take_bit_word_minus_Bit0_eq [simp]: \<open>take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1851
    = 2 * take_bit (pred_numeral n) (- numeral m)\<close> (is ?R)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1852
  and take_bit_word_minus_Bit1_eq [simp]: \<open>take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1853
    = 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\<close> (is ?S)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1854
proof -
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1855
  define w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1856
    where \<open>w = numeral m\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1857
  moreover define q :: nat
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1858
    where \<open>q = pred_numeral n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1859
  ultimately have num:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1860
    \<open>numeral m = w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1861
    \<open>numeral (num.Bit0 m) = 2 * w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1862
    \<open>numeral (num.Bit1 m) = 1 + 2 * w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1863
    \<open>numeral (Num.inc m) = 1 + w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1864
    \<open>pred_numeral n = q\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1865
    \<open>numeral n = Suc q\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1866
    by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1867
      numeral_inc numeral_eq_Suc flip: mult_2)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1868
  have even: \<open>take_bit (Suc q) (2 * w) = 2 * take_bit q w\<close> for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1869
    by (rule bit_word_eqI)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1870
      (auto simp: bit_take_bit_iff bit_double_iff)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1871
  have odd: \<open>take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\<close> for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1872
    by (rule bit_eqI)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1873
      (auto simp: bit_take_bit_iff bit_double_iff even_bit_succ_iff)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1874
  show ?P
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1875
    using even [of w] by (simp add: num)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1876
  show ?Q
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1877
    using odd [of w] by (simp add: num)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1878
  show ?R
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1879
    using even [of \<open>- w\<close>] by (simp add: num)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1880
  show ?S
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1881
    using odd [of \<open>- (1 + w)\<close>] by (simp add: num)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1882
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1883
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1884
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1885
subsection \<open>More shift operations\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1886
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1887
lift_definition signed_drop_bit :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a::len word\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1888
  is \<open>\<lambda>n. drop_bit n \<circ> signed_take_bit (LENGTH('a) - Suc 0)\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1889
  using signed_take_bit_decr_length_iff
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1890
  by (simp add: take_bit_drop_bit) force
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1891
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  1892
lemma bit_signed_drop_bit_iff [bit_simps]:
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1893
  \<open>bit (signed_drop_bit m w) n \<longleftrightarrow> bit w (if LENGTH('a) - m \<le> n \<and> n < LENGTH('a) then LENGTH('a) - 1 else m + n)\<close>
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1894
  for w :: \<open>'a::len word\<close>
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1895
  by transfer (simp add: bit_drop_bit_eq bit_signed_take_bit_iff min_def le_less not_less, auto)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1896
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  1897
lemma [code]:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  1898
  \<open>Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\<close>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  1899
  for w :: \<open>'a::len word\<close>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  1900
  by transfer simp
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  1901
73816
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1902
lemma signed_drop_bit_of_0 [simp]:
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1903
  \<open>signed_drop_bit n 0 = 0\<close>
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1904
  by transfer simp
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1905
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1906
lemma signed_drop_bit_of_minus_1 [simp]:
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1907
  \<open>signed_drop_bit n (- 1) = - 1\<close>
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1908
  by transfer simp
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
  1909
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  1910
lemma signed_drop_bit_signed_drop_bit [simp]:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  1911
  \<open>signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  1912
  for w :: \<open>'a::len word\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1913
proof (cases \<open>LENGTH('a)\<close>)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1914
  case 0
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1915
  then show ?thesis
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1916
    using len_not_eq_0 by blast
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1917
next
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1918
  case (Suc n)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1919
  then show ?thesis
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  1920
    by (force simp: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps intro!: bit_word_eqI)
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  1921
qed
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  1922
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1923
lemma signed_drop_bit_0 [simp]:
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1924
  \<open>signed_drop_bit 0 w = w\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  1925
  by transfer (simp add: take_bit_signed_take_bit)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1926
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1927
lemma sint_signed_drop_bit_eq:
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1928
  \<open>sint (signed_drop_bit n w) = drop_bit n (sint w)\<close>
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1929
  by (rule bit_eqI; cases n) (auto simp add: bit_simps not_less)
72388
633d14bd1e59 consolidated for the sake of documentation
haftmann
parents: 72292
diff changeset
  1930
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  1931
75623
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1932
subsection \<open>Single-bit operations\<close>
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1933
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1934
lemma set_bit_eq_idem_iff:
82525
02fbb93d5b01 more lemmas
haftmann
parents: 82524
diff changeset
  1935
  \<open>set_bit n w = w \<longleftrightarrow> bit w n \<or> n \<ge> LENGTH('a)\<close>
75623
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1936
  for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1937
  unfolding bit_eq_iff
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1938
  by (auto simp: bit_simps not_le)
75623
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1939
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1940
lemma unset_bit_eq_idem_iff:
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1941
  \<open>unset_bit n w = w \<longleftrightarrow> bit w n \<longrightarrow> n \<ge> LENGTH('a)\<close>
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1942
  for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1943
  unfolding bit_eq_iff
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1944
  by (auto simp: bit_simps dest: bit_imp_le_length)
75623
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1945
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1946
lemma flip_bit_eq_idem_iff:
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1947
  \<open>flip_bit n w = w \<longleftrightarrow> n \<ge> LENGTH('a)\<close>
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1948
  for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1949
  by (simp add: flip_bit_eq_if set_bit_eq_idem_iff unset_bit_eq_idem_iff)
75623
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1950
7a6301d01199 More lemmas.
haftmann
parents: 75087
diff changeset
  1951
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  1952
subsection \<open>Rotation\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  1953
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1954
lift_definition word_rotr :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1955
  is \<open>\<lambda>n k. concat_bit (LENGTH('a) - n mod LENGTH('a))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1956
    (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1957
    (take_bit (n mod LENGTH('a)) k)\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1958
  using take_bit_tightened by fastforce
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1959
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1960
lift_definition word_rotl :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1961
  is \<open>\<lambda>n k. concat_bit (n mod LENGTH('a))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1962
    (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1963
    (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1964
  using take_bit_tightened by fastforce
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1965
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1966
lift_definition word_roti :: \<open>int \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1967
  is \<open>\<lambda>r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a)))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1968
    (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1969
    (take_bit (nat (r mod int LENGTH('a))) k)\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1970
  by (smt (verit, best) len_gt_0 nat_le_iff of_nat_0_less_iff pos_mod_bound
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  1971
      take_bit_tightened)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1972
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1973
lemma word_rotl_eq_word_rotr [code]:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1974
  \<open>word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \<Rightarrow> 'a word)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1975
  by (rule ext, cases \<open>n mod LENGTH('a) = 0\<close>; transfer) simp_all
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1976
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1977
lemma word_roti_eq_word_rotr_word_rotl [code]:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1978
  \<open>word_roti i w =
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1979
    (if i \<ge> 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1980
proof (cases \<open>i \<ge> 0\<close>)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1981
  case True
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1982
  moreover define n where \<open>n = nat i\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1983
  ultimately have \<open>i = int n\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1984
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1985
  moreover have \<open>word_roti (int n) = (word_rotr n :: _ \<Rightarrow> 'a word)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1986
    by (rule ext, transfer) (simp add: nat_mod_distrib)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1987
  ultimately show ?thesis
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1988
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1989
next
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1990
  case False
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1991
  moreover define n where \<open>n = nat (- i)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1992
  ultimately have \<open>i = - int n\<close> \<open>n > 0\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1993
    by simp_all
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1994
  moreover have \<open>word_roti (- int n) = (word_rotl n :: _ \<Rightarrow> 'a word)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1995
    by (rule ext, transfer)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1996
      (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1997
  ultimately show ?thesis
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1998
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  1999
qed
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2000
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  2001
lemma bit_word_rotr_iff [bit_simps]:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2002
  \<open>bit (word_rotr m w) n \<longleftrightarrow>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2003
    n < LENGTH('a) \<and> bit w ((n + m) mod LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2004
  for w :: \<open>'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2005
proof transfer
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2006
  fix k :: int and m n :: nat
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2007
  define q where \<open>q = m mod LENGTH('a)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2008
  have \<open>q < LENGTH('a)\<close> 
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2009
    by (simp add: q_def)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2010
  then have \<open>q \<le> LENGTH('a)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2011
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2012
  have \<open>m mod LENGTH('a) = q\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2013
    by (simp add: q_def)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2014
  moreover have \<open>(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2015
    by (subst mod_add_right_eq [symmetric]) (simp add: \<open>m mod LENGTH('a) = q\<close>)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2016
  moreover have \<open>n < LENGTH('a) \<and>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2017
    bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \<longleftrightarrow>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2018
    n < LENGTH('a) \<and> bit k ((n + q) mod LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2019
    using \<open>q < LENGTH('a)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2020
    by (cases \<open>q + n \<ge> LENGTH('a)\<close>)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2021
     (auto simp: bit_concat_bit_iff bit_drop_bit_eq
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2022
        bit_take_bit_iff le_mod_geq ac_simps)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2023
  ultimately show \<open>n < LENGTH('a) \<and>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2024
    bit (concat_bit (LENGTH('a) - m mod LENGTH('a))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2025
      (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2026
      (take_bit (m mod LENGTH('a)) k)) n
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2027
    \<longleftrightarrow> n < LENGTH('a) \<and>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2028
      (n + m) mod LENGTH('a) < LENGTH('a) \<and>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2029
      bit k ((n + m) mod LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2030
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2031
qed
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2032
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  2033
lemma bit_word_rotl_iff [bit_simps]:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2034
  \<open>bit (word_rotl m w) n \<longleftrightarrow>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2035
    n < LENGTH('a) \<and> bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2036
  for w :: \<open>'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2037
  by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2038
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  2039
lemma bit_word_roti_iff [bit_simps]:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2040
  \<open>bit (word_roti k w) n \<longleftrightarrow>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2041
    n < LENGTH('a) \<and> bit w (nat ((int n + k) mod int LENGTH('a)))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2042
  for w :: \<open>'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2043
proof transfer
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2044
  fix k l :: int and n :: nat
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2045
  define m where \<open>m = nat (k mod int LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2046
  have \<open>m < LENGTH('a)\<close> 
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2047
    by (simp add: nat_less_iff m_def)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2048
  then have \<open>m \<le> LENGTH('a)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2049
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2050
  have \<open>k mod int LENGTH('a) = int m\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2051
    by (simp add: nat_less_iff m_def)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2052
  moreover have \<open>(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2053
    by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \<open>k mod int LENGTH('a) = int m\<close>)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2054
  moreover have \<open>n < LENGTH('a) \<and>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2055
    bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \<longleftrightarrow>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2056
    n < LENGTH('a) \<and> bit l ((n + m) mod LENGTH('a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2057
    using \<open>m < LENGTH('a)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2058
    by (cases \<open>m + n \<ge> LENGTH('a)\<close>)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2059
     (auto simp: bit_concat_bit_iff bit_drop_bit_eq
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2060
        bit_take_bit_iff nat_less_iff not_le not_less ac_simps
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2061
        le_diff_conv le_mod_geq)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2062
  ultimately show \<open>n < LENGTH('a)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2063
    \<and> bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a)))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2064
             (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2065
             (take_bit (nat (k mod int LENGTH('a))) l)) n \<longleftrightarrow>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2066
       n < LENGTH('a) 
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2067
    \<and> nat ((int n + k) mod int LENGTH('a)) < LENGTH('a)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2068
    \<and> bit l (nat ((int n + k) mod int LENGTH('a)))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2069
    by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2070
qed
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2071
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2072
lemma uint_word_rotr_eq:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2073
  \<open>uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2074
    (drop_bit (n mod LENGTH('a)) (uint w))
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2075
    (uint (take_bit (n mod LENGTH('a)) w))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2076
  for w :: \<open>'a::len word\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2077
  by transfer (simp add: take_bit_concat_bit_eq)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2078
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2079
lemma [code]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2080
  \<open>Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a))
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2081
    (drop_bit (n mod LENGTH('a)) (Word.the_int w))
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2082
    (Word.the_int (take_bit (n mod LENGTH('a)) w))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2083
  for w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2084
  using uint_word_rotr_eq [of n w] by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2085
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2086
    
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  2087
subsection \<open>Split and cat operations\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2088
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2089
lift_definition word_cat :: \<open>'a::len word \<Rightarrow> 'b::len word \<Rightarrow> 'c::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2090
  is \<open>\<lambda>k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2091
  by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2092
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2093
lemma word_cat_eq:
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2094
  \<open>(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2095
  for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close>
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2096
  by transfer (simp add: concat_bit_eq ac_simps)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2097
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2098
lemma word_cat_eq' [code]:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2099
  \<open>word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2100
  for a :: \<open>'a::len word\<close> and b :: \<open>'b::len word\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2101
  by transfer (simp add: concat_bit_take_bit_eq)
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2102
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  2103
lemma bit_word_cat_iff [bit_simps]:
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2104
  \<open>bit (word_cat v w :: 'c::len word) n \<longleftrightarrow> n < LENGTH('c) \<and> (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\<close> 
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2105
  for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close>
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2106
  by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff)
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  2107
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2108
definition word_split :: \<open>'a::len word \<Rightarrow> 'b::len word \<times> 'c::len word\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2109
  where \<open>word_split w =
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2110
    (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\<close>
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2111
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  2112
definition word_rcat :: \<open>'a::len word list \<Rightarrow> 'b::len word\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  2113
  where \<open>word_rcat = word_of_int \<circ> horner_sum uint (2 ^ LENGTH('a)) \<circ> rev\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  2114
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2115
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2116
subsection \<open>More on conversions\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2117
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2118
lemma int_word_sint:
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2119
  \<open>sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2120
  by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift)
46010
ebbc2d5cd720 add section headings
huffman
parents: 46009
diff changeset
  2121
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2122
lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin"
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
  2123
  by (simp add: signed_of_int)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2124
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2125
lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2126
  for w :: "'a::len word"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2127
  by transfer (simp add: take_bit_signed_take_bit)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2128
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2129
lemma bintr_uint: "LENGTH('a) \<le> n \<Longrightarrow> take_bit n (uint w) = uint w"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2130
  for w :: "'a::len word"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2131
  by transfer (simp add: min_def)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2132
46057
8664713db181 remove unnecessary intermediate lemmas
huffman
parents: 46026
diff changeset
  2133
lemma wi_bintr:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2134
  "LENGTH('a::len) \<le> n \<Longrightarrow>
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2135
    word_of_int (take_bit n w) = (word_of_int w :: 'a word)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2136
  by transfer simp
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2137
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2138
lemma word_numeral_alt: "numeral b = word_of_int (numeral b)"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2139
  by simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2140
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2141
declare word_numeral_alt [symmetric, code_abbrev]
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2142
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2143
lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2144
  by simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2145
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2146
declare word_neg_numeral_alt [symmetric, code_abbrev]
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2147
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2148
lemma uint_bintrunc [simp]:
82523
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  2149
  "uint (numeral bin :: 'a word) = take_bit LENGTH('a::len) (numeral bin)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2150
  by transfer rule
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2151
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2152
lemma uint_bintrunc_neg [simp]:
82523
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  2153
  "uint (- numeral bin :: 'a word) = take_bit LENGTH('a::len) (- numeral bin)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2154
  by transfer rule
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2155
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2156
lemma sint_sbintrunc [simp]:
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2157
  "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2158
  by transfer simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2159
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2160
lemma sint_sbintrunc_neg [simp]:
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2161
  "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2162
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2163
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2164
lemma unat_bintrunc [simp]:
82523
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  2165
  "unat (numeral bin :: 'a::len word) = take_bit LENGTH('a) (numeral bin)"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2166
  by transfer simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2167
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2168
lemma unat_bintrunc_neg [simp]:
82523
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  2169
  "unat (- numeral bin :: 'a::len word) = nat (take_bit LENGTH('a) (- numeral bin))"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2170
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2171
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2172
lemma size_0_eq: "size w = 0 \<Longrightarrow> v = w"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2173
  for v w :: "'a::len word"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2174
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2175
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2176
lemma uint_ge_0 [iff]: "0 \<le> uint x"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2177
  by (fact unsigned_greater_eq)
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2178
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2179
lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2180
  for x :: "'a::len word"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2181
  by (fact unsigned_less)
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2182
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2183
lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \<le> sint x"
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2184
  for x :: "'a::len word"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2185
  using sint_greater_eq [of x] by simp
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2186
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2187
lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)"
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2188
  for x :: "'a::len word"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2189
  using sint_less [of x] by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2190
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2191
lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2192
  for x :: "'a::len word"
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2193
  by (simp only: diff_less_0_iff_less uint_lt2p)
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2194
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2195
lemma uint_m2p_not_non_neg: "\<not> 0 \<le> uint x - 2 ^ LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2196
  for x :: "'a::len word"
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2197
  by (simp only: not_le uint_m2p_neg)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2198
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2199
lemma lt2p_lem: "LENGTH('a) \<le> n \<Longrightarrow> uint w < 2 ^ n"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2200
  for w :: "'a::len word"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2201
  using uint_bounded [of w] by (rule less_le_trans) simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2202
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2203
lemma uint_le_0_iff [simp]: "uint x \<le> 0 \<longleftrightarrow> uint x = 0"
70749
5d06b7bb9d22 More type class generalisations. Note that linorder_antisym_conv1 and linorder_antisym_conv2 no longer exist.
paulson <lp15@cam.ac.uk>
parents: 70342
diff changeset
  2204
  by (fact uint_ge_0 [THEN leD, THEN antisym_conv1])
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2205
40827
abbc05c20e24 code preprocessor setup for numerals on word type;
haftmann
parents: 39910
diff changeset
  2206
lemma uint_nat: "uint w = int (unat w)"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2207
  by transfer simp
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2208
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2209
lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2210
  by (simp flip: take_bit_eq_mod add: of_nat_take_bit)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2211
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2212
lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2213
  by (simp flip: take_bit_eq_mod add: of_nat_take_bit)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2214
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2215
lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2216
  by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2217
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2218
lemma sint_numeral:
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2219
  "sint (numeral b :: 'a::len word) =
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2220
    (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2221
  by (metis int_word_sint word_numeral_alt)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2222
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2223
lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2224
  by (fact of_int_0)
45958
c28235388c43 simplify some proofs
huffman
parents: 45957
diff changeset
  2225
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2226
lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2227
  by (fact of_int_1)
45958
c28235388c43 simplify some proofs
huffman
parents: 45957
diff changeset
  2228
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  2229
lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2230
  by simp
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  2231
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2232
lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2233
  by simp
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2234
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2235
lemma word_int_case_wi:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2236
  "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2237
  by (simp add: uint_word_of_int word_int_case_eq_uint)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2238
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2239
lemma word_int_split:
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2240
  "P (word_int_case f x) =
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2241
    (\<forall>i. x = (word_of_int i :: 'b::len word) \<and> 0 \<le> i \<and> i < 2 ^ LENGTH('b) \<longrightarrow> P (f i))"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2242
  by transfer (auto simp: take_bit_eq_mod)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2243
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2244
lemma word_int_split_asm:
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2245
  "P (word_int_case f x) =
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2246
    (\<nexists>n. x = (word_of_int n :: 'b::len word) \<and> 0 \<le> n \<and> n < 2 ^ LENGTH('b::len) \<and> \<not> P (f n))"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2247
  using word_int_split by auto
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2248
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2249
lemma uint_range_size: "0 \<le> uint w \<and> uint w < 2 ^ size w"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2250
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2251
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2252
lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \<le> sint w \<and> sint w < 2 ^ (size w - Suc 0)"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2253
  by (simp add: word_size sint_greater_eq sint_less)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2254
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2255
lemma sint_above_size: "2 ^ (size w - 1) \<le> x \<Longrightarrow> sint w < x"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2256
  for w :: "'a::len word"
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2257
  unfolding word_size by (rule less_le_trans [OF sint_lt])
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2258
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2259
lemma sint_below_size: "x \<le> - (2 ^ (size w - 1)) \<Longrightarrow> x \<le> sint w"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2260
  for w :: "'a::len word"
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  2261
  unfolding word_size by (rule order_trans [OF _ sint_ge])
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2262
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2263
lemma word_unat_eq_iff:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2264
  \<open>v = w \<longleftrightarrow> unat v = unat w\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2265
  for v w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2266
  by (fact word_eq_iff_unsigned)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2267
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2268
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  2269
subsection \<open>Testing bits\<close>
46010
ebbc2d5cd720 add section headings
huffman
parents: 46009
diff changeset
  2270
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2271
lemma bin_nth_uint_imp: "bit (uint w) n \<Longrightarrow> n < LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2272
  for w :: "'a::len word"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2273
  by (simp add: bit_uint_iff)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2274
46057
8664713db181 remove unnecessary intermediate lemmas
huffman
parents: 46026
diff changeset
  2275
lemma bin_nth_sint:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2276
  "LENGTH('a) \<le> n \<Longrightarrow>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2277
    bit (sint w) n = bit (sint w) (LENGTH('a) - 1)"
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2278
  for w :: "'a::len word"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2279
  by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2280
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2281
lemma num_of_bintr':
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2282
  "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \<Longrightarrow>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2283
    numeral a = (numeral b :: 'a word)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2284
proof (transfer fixing: a b)
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2285
  assume \<open>take_bit LENGTH('a) (numeral a :: int) = numeral b\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2286
  then have \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2287
    by simp
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2288
  then show \<open>take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2289
    by simp
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2290
qed
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2291
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2292
lemma num_of_sbintr':
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2293
  "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \<Longrightarrow>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2294
    numeral a = (numeral b :: 'a word)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2295
proof (transfer fixing: a b)
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2296
  assume \<open>signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2297
  then have \<open>take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2298
    by simp
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2299
  then show \<open>take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2300
    by (simp add: take_bit_signed_take_bit)
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2301
qed
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2302
 
46962
5bdcdb28be83 make more word theorems respect int/bin distinction
huffman
parents: 46656
diff changeset
  2303
lemma num_abs_bintr:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2304
  "(numeral x :: 'a word) =
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2305
    word_of_int (take_bit (LENGTH('a::len)) (numeral x))"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2306
  by transfer simp
46962
5bdcdb28be83 make more word theorems respect int/bin distinction
huffman
parents: 46656
diff changeset
  2307
5bdcdb28be83 make more word theorems respect int/bin distinction
huffman
parents: 46656
diff changeset
  2308
lemma num_abs_sbintr:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2309
  "(numeral x :: 'a word) =
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2310
    word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2311
  by transfer (simp add: take_bit_signed_take_bit)
46962
5bdcdb28be83 make more word theorems respect int/bin distinction
huffman
parents: 46656
diff changeset
  2312
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  2313
text \<open>
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  2314
  \<open>cast\<close> -- note, no arg for new length, as it's determined by type of result,
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  2315
  thus in \<open>cast w = w\<close>, the type means cast to length of \<open>w\<close>!
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  2316
\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2317
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  2318
lemma bit_ucast_iff:
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2319
  \<open>bit (ucast a :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a::len) \<and> bit a n\<close>
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2320
  by transfer (simp add: bit_take_bit_iff)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2321
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2322
lemma ucast_id [simp]: "ucast w = w"
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2323
  by transfer simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2324
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2325
lemma scast_id [simp]: "scast w = w"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2326
  by transfer (simp add: take_bit_signed_take_bit)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2327
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  2328
lemma ucast_mask_eq:
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  2329
  \<open>ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2330
  by (simp add: bit_eq_iff) (auto simp: bit_mask_iff bit_ucast_iff)
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  2331
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  2332
\<comment> \<open>literal u(s)cast\<close>
46001
0b562d564d5f redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents: 46000
diff changeset
  2333
lemma ucast_bintr [simp]:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2334
  "ucast (numeral w :: 'a::len word) =
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2335
    word_of_int (take_bit (LENGTH('a)) (numeral w))"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2336
  by transfer simp
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2337
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2338
(* TODO: neg_numeral *)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2339
46001
0b562d564d5f redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents: 46000
diff changeset
  2340
lemma scast_sbintr [simp]:
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2341
  "scast (numeral w ::'a::len word) =
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2342
    word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2343
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2344
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2345
lemma source_size: "source_size (c::'a::len word \<Rightarrow> _) = LENGTH('a)"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2346
  by transfer simp
46011
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  2347
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2348
lemma target_size: "target_size (c::_ \<Rightarrow> 'b::len word) = LENGTH('b)"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2349
  by transfer simp
46011
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  2350
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2351
lemma is_down: "is_down c \<longleftrightarrow> LENGTH('b) \<le> LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2352
  for c :: "'a::len word \<Rightarrow> 'b::len word"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2353
  by transfer simp
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2354
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2355
lemma is_up: "is_up c \<longleftrightarrow> LENGTH('a) \<le> LENGTH('b)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2356
  for c :: "'a::len word \<Rightarrow> 'b::len word"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2357
  by transfer simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2358
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2359
lemma is_up_down:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2360
  \<open>is_up c \<longleftrightarrow> is_down d\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2361
  for c :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2362
  and d :: \<open>'b::len word \<Rightarrow> 'a::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2363
  by transfer simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2364
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2365
context
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2366
  fixes dummy_types :: \<open>'a::len \<times> 'b::len\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2367
begin
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2368
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2369
private abbreviation (input) UCAST :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2370
  where \<open>UCAST == ucast\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2371
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2372
private abbreviation (input) SCAST :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2373
  where \<open>SCAST == scast\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2374
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2375
lemma down_cast_same:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2376
  \<open>UCAST = scast\<close> if \<open>is_down UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2377
  by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2378
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2379
lemma sint_up_scast:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2380
  \<open>sint (SCAST w) = sint w\<close> if \<open>is_up SCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2381
  using that by transfer (simp add: min_def Suc_leI le_diff_iff)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2382
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2383
lemma uint_up_ucast:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2384
  \<open>uint (UCAST w) = uint w\<close> if \<open>is_up UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2385
  using that by transfer (simp add: min_def)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2386
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2387
lemma ucast_up_ucast:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2388
  \<open>ucast (UCAST w) = ucast w\<close> if \<open>is_up UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2389
  using that by transfer (simp add: ac_simps)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2390
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2391
lemma ucast_up_ucast_id:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2392
  \<open>ucast (UCAST w) = w\<close> if \<open>is_up UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2393
  using that by (simp add: ucast_up_ucast)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2394
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2395
lemma scast_up_scast:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2396
  \<open>scast (SCAST w) = scast w\<close> if \<open>is_up SCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2397
  using that by transfer (simp add: ac_simps)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2398
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2399
lemma scast_up_scast_id:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2400
  \<open>scast (SCAST w) = w\<close> if \<open>is_up SCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2401
  using that by (simp add: scast_up_scast)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2402
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2403
lemma isduu:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2404
  \<open>is_up UCAST\<close> if \<open>is_down d\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2405
    for d :: \<open>'b word \<Rightarrow> 'a word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2406
  using that is_up_down [of UCAST d] by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2407
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2408
lemma isdus:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2409
  \<open>is_up SCAST\<close> if \<open>is_down d\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2410
    for d :: \<open>'b word \<Rightarrow> 'a word\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2411
  using that is_up_down [of SCAST d] by simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2412
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2413
lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id]
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2414
lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2415
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2416
lemma up_ucast_surj:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2417
  \<open>surj (ucast :: 'b word \<Rightarrow> 'a word)\<close> if \<open>is_up UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2418
  by (rule surjI) (use that in \<open>rule ucast_up_ucast_id\<close>)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2419
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2420
lemma up_scast_surj:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2421
  \<open>surj (scast :: 'b word \<Rightarrow> 'a word)\<close> if \<open>is_up SCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2422
  by (rule surjI) (use that in \<open>rule scast_up_scast_id\<close>)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2423
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2424
lemma down_ucast_inj:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2425
  \<open>inj_on UCAST A\<close> if \<open>is_down (ucast :: 'b word \<Rightarrow> 'a word)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2426
  by (rule inj_on_inverseI) (use that in \<open>rule ucast_down_ucast_id\<close>)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2427
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2428
lemma down_scast_inj:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2429
  \<open>inj_on SCAST A\<close> if \<open>is_down (scast :: 'b word \<Rightarrow> 'a word)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2430
  by (rule inj_on_inverseI) (use that in \<open>rule scast_down_scast_id\<close>)
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2431
  
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2432
lemma ucast_down_wi:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2433
  \<open>UCAST (word_of_int x) = word_of_int x\<close> if \<open>is_down UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2434
  using that by transfer simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2435
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2436
lemma ucast_down_no:
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2437
  \<open>UCAST (numeral bin) = numeral bin\<close> if \<open>is_down UCAST\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2438
  using that by transfer simp
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2439
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2440
end
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2441
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2442
lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2443
72000
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2444
lemma bit_last_iff:
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2445
  \<open>bit w (LENGTH('a) - Suc 0) \<longleftrightarrow> sint w < 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2446
  for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2447
  by (simp add: bit_unsigned_iff sint_uint)
72000
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2448
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2449
lemma drop_bit_eq_zero_iff_not_bit_last:
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2450
  \<open>drop_bit (LENGTH('a) - Suc 0) w = 0 \<longleftrightarrow> \<not> bit w (LENGTH('a) - Suc 0)\<close>
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2451
  for w :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2452
proof (cases \<open>LENGTH('a)\<close>)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2453
  case (Suc n)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2454
  then show ?thesis
72000
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2455
    apply transfer
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2456
    apply (simp add: take_bit_drop_bit)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2457
    by (simp add: bit_iff_odd_drop_bit drop_bit_take_bit odd_iff_mod_2_eq_one)
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2458
qed auto
72000
379d0c207c29 separation of traditional bit operations
haftmann
parents: 71997
diff changeset
  2459
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2460
lemma unat_div:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2461
  \<open>unat (x div y) = unat x div unat y\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2462
  by (fact unat_div_distrib)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2463
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2464
lemma unat_mod:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2465
  \<open>unat (x mod y) = unat x mod unat y\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2466
  by (fact unat_mod_distrib)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2467
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2468
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  2469
subsection \<open>Word Arithmetic\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2470
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2471
lemmas less_eq_word_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2472
  word_le_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2473
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2474
lemmas less_word_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2475
  word_less_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2476
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2477
lemmas less_eq_word_minus_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2478
  word_le_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2479
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2480
lemmas less_word_minus_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2481
  word_less_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2482
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2483
lemmas less_eq_word_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2484
  word_le_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2485
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2486
lemmas less_word_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2487
  word_less_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2488
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2489
lemmas less_eq_word_minus_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2490
  word_le_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2491
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2492
lemmas less_word_minus_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2493
  word_less_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2494
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2495
lemmas less_word_numeral_minus_1 [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2496
  word_less_def [of \<open>numeral a\<close> \<open>- 1\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2497
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2498
lemmas less_word_minus_numeral_minus_1 [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2499
  word_less_def [of \<open>- numeral a\<close> \<open>- 1\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2500
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2501
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2502
lemmas sless_eq_word_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2503
  word_sle_eq [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2504
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2505
lemmas sless_word_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2506
  word_sless_alt [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2507
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2508
lemmas sless_eq_word_minus_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2509
  word_sle_eq [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2510
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2511
lemmas sless_word_minus_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2512
  word_sless_alt [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2513
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2514
lemmas sless_eq_word_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2515
  word_sle_eq [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2516
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2517
lemmas sless_word_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2518
  word_sless_alt [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2519
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2520
lemmas sless_eq_word_minus_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2521
  word_sle_eq [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2522
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2523
lemmas sless_word_minus_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2524
  word_sless_alt [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2525
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2526
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2527
lemmas div_word_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2528
  word_div_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2529
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2530
lemmas div_word_minus_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2531
  word_div_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2532
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2533
lemmas div_word_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2534
  word_div_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2535
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2536
lemmas div_word_minus_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2537
  word_div_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2538
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2539
lemmas div_word_minus_1_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2540
  word_div_def [of \<open>- 1\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2541
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2542
lemmas div_word_minus_1_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2543
  word_div_def [of \<open>- 1\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2544
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2545
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2546
lemmas mod_word_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2547
  word_mod_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2548
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2549
lemmas mod_word_minus_numeral_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2550
  word_mod_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2551
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2552
lemmas mod_word_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2553
  word_mod_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2554
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2555
lemmas mod_word_minus_numeral_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2556
  word_mod_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2557
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2558
lemmas mod_word_minus_1_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2559
  word_mod_def [of \<open>- 1\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2560
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2561
lemmas mod_word_minus_1_minus_numeral [simp] =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2562
  word_mod_def [of \<open>- 1\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2563
  for a b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2564
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2565
lemma signed_drop_bit_of_1 [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2566
  \<open>signed_drop_bit n (1 :: 'a::len word) = of_bool (LENGTH('a) = 1 \<or> n = 0)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2567
  apply (transfer fixing: n)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2568
  apply (cases \<open>LENGTH('a)\<close>)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2569
   apply (auto simp: take_bit_signed_take_bit)
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2570
  apply (auto simp: take_bit_drop_bit gr0_conv_Suc simp flip: take_bit_eq_self_iff_drop_bit_eq_0)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2571
  done
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2572
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2573
lemma take_bit_word_beyond_length_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2574
  \<open>take_bit n w = w\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2575
  by (simp add: take_bit_word_eq_self that)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2576
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2577
lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2578
lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2579
lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  2580
lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2581
lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2582
lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2583
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2584
lemma size_0_same': "size w = 0 \<Longrightarrow> w = v"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2585
  for v w :: "'a::len word"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2586
  by (unfold word_size) simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2587
45816
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  2588
lemmas size_0_same = size_0_same' [unfolded word_size]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2589
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2590
lemmas unat_eq_0 = unat_0_iff
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2591
lemmas unat_eq_zero = unat_0_iff
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2592
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2593
lemma mask_1: "mask 1 = 1"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2594
  by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2595
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2596
lemma mask_Suc_0: "mask (Suc 0) = 1"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2597
  by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2598
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2599
lemma bin_last_bintrunc: "odd (take_bit l n) \<longleftrightarrow> l > 0 \<and> odd n"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2600
  by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2601
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2602
lemma push_bit_word_beyond [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2603
  \<open>push_bit n w = 0\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2604
  using that by (transfer fixing: n) (simp add: take_bit_push_bit)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2605
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2606
lemma drop_bit_word_beyond [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2607
  \<open>drop_bit n w = 0\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2608
  using that by (transfer fixing: n) (simp add: drop_bit_take_bit)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2609
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2610
lemma signed_drop_bit_beyond:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2611
  \<open>signed_drop_bit n w = (if bit w (LENGTH('a) - Suc 0) then - 1 else 0)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2612
  if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2613
  by (rule bit_word_eqI) (simp add: bit_signed_drop_bit_iff that)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2614
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2615
lemma take_bit_numeral_minus_numeral_word [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2616
  \<open>take_bit (numeral m) (- numeral n :: 'a::len word) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2617
    (case take_bit_num (numeral m) n of None \<Rightarrow> 0 | Some q \<Rightarrow> take_bit (numeral m) (2 ^ numeral m - numeral q))\<close> (is \<open>?lhs = ?rhs\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2618
proof (cases \<open>LENGTH('a) \<le> numeral m\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2619
  case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2620
  then have *: \<open>(take_bit (numeral m) :: 'a word \<Rightarrow> 'a word) = id\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2621
    by (simp add: fun_eq_iff take_bit_word_eq_self)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2622
  have **: \<open>2 ^ numeral m = (0 :: 'a word)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2623
    using True by (simp flip: exp_eq_zero_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2624
  show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2625
    by (auto simp only: * ** split: option.split
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2626
      dest!: take_bit_num_eq_None_imp [where ?'a = \<open>'a word\<close>] take_bit_num_eq_Some_imp [where ?'a = \<open>'a word\<close>])
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2627
      simp_all
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2628
next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2629
  case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2630
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2631
    by (transfer fixing: m n) simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2632
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2633
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2634
lemma of_nat_inverse:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2635
  \<open>word_of_nat r = a \<Longrightarrow> r < 2 ^ LENGTH('a) \<Longrightarrow> unat a = r\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2636
  for a :: \<open>'a::len word\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2637
  by (metis id_apply of_nat_eq_id take_bit_nat_eq_self_iff unsigned_of_nat)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2638
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2639
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  2640
subsection \<open>Transferring goals from words to ints\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2641
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2642
lemma word_ths:
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2643
  shows word_succ_p1: "word_succ a = a + 1"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2644
    and word_pred_m1: "word_pred a = a - 1"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2645
    and word_pred_succ: "word_pred (word_succ a) = a"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2646
    and word_succ_pred: "word_succ (word_pred a) = a"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2647
    and word_mult_succ: "word_succ a * b = b + a * b"
47374
9475d524bafb set up and use lift_definition for word operations
huffman
parents: 47372
diff changeset
  2648
  by (transfer, simp add: algebra_simps)+
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2649
45816
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  2650
lemma uint_cong: "x = y \<Longrightarrow> uint x = uint y"
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  2651
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2652
55818
d8b2f50705d0 more precise imports;
haftmann
parents: 55817
diff changeset
  2653
lemma uint_word_ariths:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2654
  fixes a b :: "'a::len word"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2655
  shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)"
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2656
    and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)"
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2657
    and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)"
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2658
    and "uint (- a) = - uint a mod 2 ^ LENGTH('a)"
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2659
    and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)"
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2660
    and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)"
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2661
    and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)"
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2662
    and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2663
  by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod)
55818
d8b2f50705d0 more precise imports;
haftmann
parents: 55817
diff changeset
  2664
d8b2f50705d0 more precise imports;
haftmann
parents: 55817
diff changeset
  2665
lemma uint_word_arith_bintrs:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2666
  fixes a b :: "'a::len word"
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2667
  shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2668
    and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2669
    and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2670
    and "uint (- a) = take_bit (LENGTH('a)) (- uint a)"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2671
    and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2672
    and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2673
    and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2674
    and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1"
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  2675
  by (simp_all add: uint_word_ariths take_bit_eq_mod)
55818
d8b2f50705d0 more precise imports;
haftmann
parents: 55817
diff changeset
  2676
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2677
context
55818
d8b2f50705d0 more precise imports;
haftmann
parents: 55817
diff changeset
  2678
  fixes a b :: "'a::len word"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2679
begin
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2680
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2681
lemma sint_word_add: "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2682
  by transfer (simp add: signed_take_bit_add)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2683
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2684
lemma sint_word_diff: "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2685
  by transfer (simp add: signed_take_bit_diff)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2686
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2687
lemma sint_word_mult: "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2688
  by transfer (simp add: signed_take_bit_mult)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2689
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2690
lemma sint_word_minus: "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2691
  by transfer (simp add: signed_take_bit_minus)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2692
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2693
lemma sint_word_succ: "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2694
  by (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2695
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2696
lemma sint_word_pred: "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2697
  by (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2698
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2699
lemma sint_word_01:
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2700
  "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2701
  "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2702
  by (simp_all add: sint_uint)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2703
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2704
end
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2705
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2706
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2707
lemmas sint_word_ariths = 
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2708
  sint_word_add sint_word_diff sint_word_mult sint_word_minus 
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2709
  sint_word_succ sint_word_pred sint_word_01
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  2710
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58061
diff changeset
  2711
lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)"
47374
9475d524bafb set up and use lift_definition for word operations
huffman
parents: 47372
diff changeset
  2712
  unfolding word_pred_m1 by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2713
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2714
lemma succ_pred_no [simp]:
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2715
    "word_succ (numeral w) = numeral w + 1"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2716
    "word_pred (numeral w) = numeral w - 1"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2717
    "word_succ (- numeral w) = - numeral w + 1"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2718
    "word_pred (- numeral w) = - numeral w - 1"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2719
  by (simp_all add: word_succ_p1 word_pred_m1)
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2720
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2721
lemma word_sp_01 [simp]:
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2722
  "word_succ (- 1) = 0 \<and> word_succ 0 = 1 \<and> word_pred 0 = - 1 \<and> word_pred 1 = 0"
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2723
  by (simp_all add: word_succ_p1 word_pred_m1)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2724
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  2725
\<comment> \<open>alternative approach to lifting arithmetic equalities\<close>
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2726
lemma word_of_int_Ex: "\<exists>y. x = word_of_int y"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2727
  by (rule_tac x="uint x" in exI) simp
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2728
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2729
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  2730
subsection \<open>Order on fixed-length words\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2731
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2732
lift_definition udvd :: \<open>'a::len word \<Rightarrow> 'a::len word \<Rightarrow> bool\<close> (infixl \<open>udvd\<close> 50)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2733
  is \<open>\<lambda>k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\<close> by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2734
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2735
lemma udvd_iff_dvd:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2736
  \<open>x udvd y \<longleftrightarrow> unat x dvd unat y\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2737
  by transfer (simp add: nat_dvd_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2738
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2739
lemma udvd_iff_dvd_int:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2740
  \<open>v udvd w \<longleftrightarrow> uint v dvd uint w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2741
  by transfer rule
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2742
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2743
lemma udvdI [intro]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2744
  \<open>v udvd w\<close> if \<open>unat w = unat v * unat u\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2745
proof -
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2746
  from that have \<open>unat v dvd unat w\<close> ..
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2747
  then show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2748
    by (simp add: udvd_iff_dvd)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2749
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2750
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2751
lemma udvdE [elim]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2752
  fixes v w :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2753
  assumes \<open>v udvd w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2754
  obtains u :: \<open>'a word\<close> where \<open>unat w = unat v * unat u\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2755
proof (cases \<open>v = 0\<close>)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2756
  case True
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2757
  moreover from True \<open>v udvd w\<close> have \<open>w = 0\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2758
    by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2759
  ultimately show thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2760
    using that by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2761
next
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2762
  case False
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2763
  then have \<open>unat v > 0\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2764
    by (simp add: unat_gt_0)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2765
  from \<open>v udvd w\<close> have \<open>unat v dvd unat w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2766
    by (simp add: udvd_iff_dvd)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2767
  then obtain n where \<open>unat w = unat v * n\<close> ..
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2768
  moreover have \<open>n < 2 ^ LENGTH('a)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2769
  proof (rule ccontr)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2770
    assume \<open>\<not> n < 2 ^ LENGTH('a)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2771
    then have \<open>n \<ge> 2 ^ LENGTH('a)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2772
      by (simp add: not_le)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2773
    then have \<open>unat v * n \<ge> 2 ^ LENGTH('a)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2774
      using \<open>unat v > 0\<close> mult_le_mono [of 1 \<open>unat v\<close> \<open>2 ^ LENGTH('a)\<close> n]
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2775
      by simp
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2776
    with \<open>unat w = unat v * n\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2777
    have \<open>unat w \<ge> 2 ^ LENGTH('a)\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2778
      by simp
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2779
    with unsigned_less [of w, where ?'a = nat] show False
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2780
      by linarith
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2781
  qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2782
  ultimately have \<open>unat w = unat v * unat (word_of_nat n :: 'a word)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2783
    by (auto simp: take_bit_nat_eq_self_iff unsigned_of_nat intro: sym)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2784
  with that show thesis .
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2785
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2786
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2787
lemma udvd_imp_mod_eq_0:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2788
  \<open>w mod v = 0\<close> if \<open>v udvd w\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2789
  using that by transfer simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2790
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2791
lemma mod_eq_0_imp_udvd [intro?]:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2792
  \<open>v udvd w\<close> if \<open>w mod v = 0\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2793
  by (metis mod_0_imp_dvd that udvd_iff_dvd unat_0 unat_mod_distrib)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2794
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2795
lemma udvd_imp_dvd:
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2796
  \<open>v dvd w\<close> if \<open>v udvd w\<close> for v w :: \<open>'a::len word\<close>
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2797
proof -
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2798
  from that obtain u :: \<open>'a word\<close> where \<open>unat w = unat v * unat u\<close> ..
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2799
  then have \<open>w = v * u\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2800
    by (metis of_nat_mult of_nat_unat word_mult_def word_of_int_uint)
72280
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2801
  then show \<open>v dvd w\<close> ..
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2802
qed
db43ee05066d canonical enum instance for word
haftmann
parents: 72265
diff changeset
  2803
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2804
lemma exp_dvd_iff_exp_udvd:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2805
  \<open>2 ^ n dvd w \<longleftrightarrow> 2 ^ n udvd w\<close> for v w :: \<open>'a::len word\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2806
proof
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2807
  assume \<open>2 ^ n udvd w\<close> then show \<open>2 ^ n dvd w\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2808
    by (rule udvd_imp_dvd) 
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2809
next
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2810
  assume \<open>2 ^ n dvd w\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2811
  then obtain u :: \<open>'a word\<close> where \<open>w = 2 ^ n * u\<close> ..
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2812
  then have \<open>w = push_bit n u\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2813
    by (simp add: push_bit_eq_mult)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2814
  then show \<open>2 ^ n udvd w\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2815
    by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2816
qed
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72280
diff changeset
  2817
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2818
lemma udvd_nat_alt:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2819
  \<open>a udvd b \<longleftrightarrow> (\<exists>n. unat b = n * unat a)\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2820
  by (auto simp: udvd_iff_dvd)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2821
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2822
lemma udvd_unfold_int:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  2823
  \<open>a udvd b \<longleftrightarrow> (\<exists>n\<ge>0. uint b = n * uint a)\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2824
  unfolding udvd_iff_dvd_int
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2825
  by (metis dvd_div_mult_self dvd_triv_right uint_div_distrib uint_ge_0)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2826
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2827
lemma unat_minus_one:
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2828
  \<open>unat (w - 1) = unat w - 1\<close> if \<open>w \<noteq> 0\<close>
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2829
proof -
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2830
  have "0 \<le> uint w" by (fact uint_nonnegative)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2831
  moreover from that have "0 \<noteq> uint w"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2832
    by (simp add: uint_0_iff)
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2833
  ultimately have "1 \<le> uint w"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2834
    by arith
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2835
  from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2836
    by arith
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2837
  with \<open>1 \<le> uint w\<close> have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1"
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2838
    by (auto intro: mod_pos_pos_trivial)
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2839
  with \<open>1 \<le> uint w\<close> have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  2840
    by (auto simp del: nat_uint_eq)
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2841
  then show ?thesis
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2842
    by (metis uint_word_ariths(6) unat_eq_nat_uint word_pred_m1)
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2843
qed
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2844
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2845
lemma measure_unat: "p \<noteq> 0 \<Longrightarrow> unat (p - 1) < unat p"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2846
  by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric])
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2847
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  2848
lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0]
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  2849
lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2850
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2851
lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2852
  for x :: "'a::len word" and y :: "'b::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2853
  using uint_ge_0 [of y] uint_lt2p [of x] by arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2854
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2855
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  2856
subsection \<open>Conditions for the addition (etc) of two words to overflow\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2857
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2858
lemma uint_add_lem:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2859
  "(uint x + uint y < 2 ^ LENGTH('a)) =
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2860
    (uint (x + y) = uint x + uint y)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2861
  for x y :: "'a::len word"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  2862
  by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1))
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2863
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2864
lemma uint_mult_lem:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2865
  "(uint x * uint y < 2 ^ LENGTH('a)) =
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2866
    (uint (x * y) = uint x * uint y)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2867
  for x y :: "'a::len word"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  2868
  by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3))
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2869
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2870
lemma uint_sub_lem: "uint x \<ge> uint y \<longleftrightarrow> uint (x - y) = uint x - uint y"
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  2871
  by (simp add: uint_word_arith_bintrs take_bit_int_eq_self_iff)
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2872
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2873
lemma uint_add_le: "uint (x + y) \<le> uint x + uint y"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  2874
  unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) 
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2875
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2876
lemma uint_sub_ge: "uint (x - y) \<ge> uint x - uint y"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2877
  by (smt (verit, ccfv_SIG) uint_nonnegative uint_sub_lem)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2878
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2879
lemma int_mod_ge: \<open>a \<le> a mod n\<close> if \<open>a < n\<close> \<open>0 < n\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2880
  for a n :: int
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 75623
diff changeset
  2881
  using that order.trans [of a 0 \<open>a mod n\<close>] by (cases \<open>a < 0\<close>) auto
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  2882
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2883
lemma mod_add_if_z:
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2884
  "\<lbrakk>x < z; y < z; 0 \<le> y; 0 \<le> x; 0 \<le> z\<rbrakk> \<Longrightarrow>
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2885
    (x + y) mod z = (if x + y < z then x + y else x + y - z)"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2886
  for x y z :: int
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2887
  by (smt (verit, best) minus_mod_self2 mod_pos_pos_trivial)
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2888
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2889
lemma uint_plus_if':
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2890
  "uint (a + b) =
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2891
    (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2892
     else uint a + uint b - 2 ^ LENGTH('a))"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2893
  for a b :: "'a::len word"
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2894
  using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths)
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2895
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2896
lemma mod_sub_if_z:
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2897
  "\<lbrakk>x < z; y < z; 0 \<le> y; 0 \<le> x; 0 \<le> z\<rbrakk> \<Longrightarrow>
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2898
    (x - y) mod z = (if y \<le> x then x - y else x - y + z)"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2899
  for x y z :: int
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2900
  using mod_pos_pos_trivial [of "x - y + z" z] by (auto simp: not_le)
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2901
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2902
lemma uint_sub_if':
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2903
  "uint (a - b) =
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  2904
    (if uint b \<le> uint a then uint a - uint b
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2905
     else uint a - uint b + 2 ^ LENGTH('a))"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2906
  for a b :: "'a::len word"
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2907
  using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths)
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  2908
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2909
lemma word_of_int_inverse:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2910
  "word_of_int r = a \<Longrightarrow> 0 \<le> r \<Longrightarrow> r < 2 ^ LENGTH('a) \<Longrightarrow> uint a = r"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2911
  for a :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  2912
  by transfer (simp add: take_bit_int_eq_self)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2913
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2914
lemma unat_split: "P (unat x) \<longleftrightarrow> (\<forall>n. of_nat n = x \<and> n < 2^LENGTH('a) \<longrightarrow> P n)"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2915
  for x :: "'a::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2916
  by (auto simp: unsigned_of_nat take_bit_nat_eq_self)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2917
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2918
lemma unat_split_asm: "P (unat x) \<longleftrightarrow> (\<nexists>n. of_nat n = x \<and> n < 2^LENGTH('a) \<and> \<not> P n)"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2919
  for x :: "'a::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2920
  using unat_split by auto
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2921
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2922
lemma un_ui_le:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2923
  \<open>unat a \<le> unat b \<longleftrightarrow> uint a \<le> uint b\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2924
  by transfer (simp add: nat_le_iff) 
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2925
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2926
lemma unat_plus_if':
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2927
  \<open>unat (a + b) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2928
    (if unat a + unat b < 2 ^ LENGTH('a)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2929
    then unat a + unat b
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2930
    else unat a + unat b - 2 ^ LENGTH('a))\<close> for a b :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2931
  apply (auto simp: not_less le_iff_add)
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2932
  using of_nat_inverse apply force
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2933
  by (smt (verit, ccfv_SIG) numeral_Bit0 numerals(1) of_nat_0_le_iff of_nat_1 of_nat_add of_nat_eq_iff of_nat_power of_nat_unat uint_plus_if')
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2934
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2935
lemma unat_sub_if_size:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2936
  "unat (x - y) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2937
    (if unat y \<le> unat x
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2938
     then unat x - unat y
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2939
     else unat x + 2 ^ size x - unat y)"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2940
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2941
  { assume xy: "\<not> uint y \<le> uint x"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2942
    have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x + 2 ^ LENGTH('a) - uint y)"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2943
      by simp
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2944
    also have "\<dots> = nat (uint x + 2 ^ LENGTH('a)) - nat (uint y)"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2945
      by (simp add: nat_diff_distrib')
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2946
    also have "\<dots> = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2947
      by (simp add: nat_add_distrib nat_power_eq)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2948
    finally have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)" .
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2949
  }
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2950
  then show ?thesis
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2951
    by (metis nat_diff_distrib' uint_range_size uint_sub_if' un_ui_le unat_eq_nat_uint
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  2952
        word_size)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2953
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2954
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2955
lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2956
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2957
lemma uint_split:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2958
  "P (uint x) = (\<forall>i. word_of_int i = x \<and> 0 \<le> i \<and> i < 2^LENGTH('a) \<longrightarrow> P i)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2959
  for x :: "'a::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2960
  by transfer (auto simp: take_bit_eq_mod)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2961
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2962
lemma uint_split_asm:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  2963
  "P (uint x) = (\<nexists>i. word_of_int i = x \<and> 0 \<le> i \<and> i < 2^LENGTH('a) \<and> \<not> P i)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  2964
  for x :: "'a::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  2965
  by (auto simp: unsigned_of_int take_bit_int_eq_self)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2966
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2967
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2968
subsection \<open>Some proof tool support\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2969
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2970
\<comment> \<open>use this to stop, eg. \<open>2 ^ LENGTH(32)\<close> being simplified\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2971
lemma power_False_cong: "False \<Longrightarrow> a ^ b = c ^ d"
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2972
  by auto
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2973
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2974
lemmas unat_splits = unat_split unat_split_asm
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2975
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2976
lemmas unat_arith_simps =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2977
  word_le_nat_alt word_less_nat_alt
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2978
  word_unat_eq_iff
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2979
  unat_sub_if' unat_plus_if' unat_div unat_mod
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2980
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2981
lemmas uint_splits = uint_split uint_split_asm
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2982
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  2983
lemmas uint_arith_simps =
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2984
  word_le_def word_less_alt
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  2985
  word_uint_eq_iff
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2986
  uint_sub_if' uint_plus_if'
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  2987
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2988
\<comment> \<open>\<open>unat_arith_tac\<close>: tactic to reduce word arithmetic to \<open>nat\<close>, try to solve via \<open>arith\<close>\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2989
ML \<open>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2990
val unat_arith_simpset =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2991
  @{context} (* TODO: completely explicitly determined simpset *)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2992
  |> fold Simplifier.add_simp @{thms unat_arith_simps}
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2993
  |> fold Splitter.add_split @{thms if_split_asm}
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2994
  |> fold Simplifier.add_cong @{thms power_False_cong}
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2995
  |> simpset_of
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2996
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2997
fun unat_arith_tacs ctxt =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2998
  let
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2999
    fun arith_tac' n t =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3000
      Arith_Data.arith_tac ctxt n t
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3001
        handle Cooper.COOPER _ => Seq.empty;
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3002
  in
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3003
    [ clarify_tac ctxt 1,
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3004
      full_simp_tac (put_simpset unat_arith_simpset ctxt) 1,
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3005
      ALLGOALS (full_simp_tac
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3006
        (put_simpset HOL_ss ctxt
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3007
          |> fold Splitter.add_split @{thms unat_splits}
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3008
          |> fold Simplifier.add_cong @{thms power_False_cong})),
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3009
      rewrite_goals_tac ctxt @{thms word_size},
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3010
      ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3011
                         REPEAT (eresolve_tac ctxt [conjE] n) THEN
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3012
                         REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)),
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3013
      TRYALL arith_tac' ]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3014
  end
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3015
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3016
fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt))
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3017
\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3018
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3019
method_setup unat_arith =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3020
  \<open>Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3021
  "solving word arithmetic via natural numbers and arith"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3022
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  3023
\<comment> \<open>\<open>uint_arith_tac\<close>: reduce to arithmetic on int, try to solve by arith\<close>
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3024
ML \<open>
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3025
val uint_arith_simpset =
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3026
  @{context} (* TODO: completely explicitly determined simpset *)
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3027
  |> fold Simplifier.add_simp @{thms uint_arith_simps}
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3028
  |> fold Splitter.add_split @{thms if_split_asm}
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3029
  |> fold Simplifier.add_cong @{thms power_False_cong}
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3030
  |> simpset_of;
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3031
  
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3032
fun uint_arith_tacs ctxt =
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3033
  let
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3034
    fun arith_tac' n t =
59657
2441a80fb6c1 eliminated unused arith "verbose" flag -- tools that need options can use the context;
wenzelm
parents: 59498
diff changeset
  3035
      Arith_Data.arith_tac ctxt n t
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3036
        handle Cooper.COOPER _ => Seq.empty;
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3037
  in
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 41550
diff changeset
  3038
    [ clarify_tac ctxt 1,
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3039
      full_simp_tac (put_simpset uint_arith_simpset ctxt) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51375
diff changeset
  3040
      ALLGOALS (full_simp_tac
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51375
diff changeset
  3041
        (put_simpset HOL_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51375
diff changeset
  3042
          |> fold Splitter.add_split @{thms uint_splits}
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51375
diff changeset
  3043
          |> fold Simplifier.add_cong @{thms power_False_cong})),
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3044
      rewrite_goals_tac ctxt @{thms word_size},
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59487
diff changeset
  3045
      ALLGOALS  (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60429
diff changeset
  3046
                         REPEAT (eresolve_tac ctxt [conjE] n) THEN
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3047
                         REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3048
                                 THEN assume_tac ctxt n
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58874
diff changeset
  3049
                                 THEN assume_tac ctxt n)),
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3050
      TRYALL arith_tac' ]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3051
  end
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3052
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3053
fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt))
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3054
\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3055
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3056
method_setup uint_arith =
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3057
  \<open>Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3058
  "solving word arithmetic via integers and arith"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3059
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3060
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3061
subsection \<open>More on overflows and monotonicity\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3062
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3063
lemma no_plus_overflow_uint_size: "x \<le> x + y \<longleftrightarrow> uint x + uint y < 2 ^ size x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3064
  for x y :: "'a::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3065
  by (auto simp: word_size word_le_def uint_add_lem uint_sub_lem)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3066
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3067
lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3068
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3069
lemma no_ulen_sub: "x \<ge> x - y \<longleftrightarrow> uint y \<le> uint x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3070
  for x y :: "'a::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3071
  by (auto simp: word_size word_le_def uint_add_lem uint_sub_lem)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3072
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3073
lemma no_olen_add': "x \<le> y + x \<longleftrightarrow> uint y + uint x < 2 ^ LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3074
  for x y :: "'a::len word"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  3075
  by (simp add: ac_simps no_olen_add)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3076
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3077
lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]]
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3078
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3079
lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem]
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3080
lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1]
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3081
lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3082
lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3083
lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def]
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3084
lemmas word_sub_le = word_sub_le_iff [THEN iffD2]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3085
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3086
lemma word_less_sub1: "x \<noteq> 0 \<Longrightarrow> 1 < x \<longleftrightarrow> 0 < x - 1"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3087
  for x :: "'a::len word"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3088
  by transfer (simp add: take_bit_decr_eq) 
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3089
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3090
lemma word_le_sub1: "x \<noteq> 0 \<Longrightarrow> 1 \<le> x \<longleftrightarrow> 0 \<le> x - 1"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3091
  for x :: "'a::len word"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3092
  by transfer (simp add: int_one_le_iff_zero_less less_le)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3093
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3094
lemma sub_wrap_lt: "x < x - z \<longleftrightarrow> x < z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3095
  for x z :: "'a::len word"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3096
  by (meson linorder_not_le word_sub_le_iff)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3097
  
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3098
lemma sub_wrap: "x \<le> x - z \<longleftrightarrow> z = 0 \<or> x < z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3099
  for x z :: "'a::len word"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3100
  by (simp add: le_less sub_wrap_lt ac_simps)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3101
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3102
lemma plus_minus_not_NULL_ab: "x \<le> ab - c \<Longrightarrow> c \<le> ab \<Longrightarrow> c \<noteq> 0 \<Longrightarrow> x + c \<noteq> 0"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3103
  for x ab c :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3104
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3105
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3106
lemma plus_minus_no_overflow_ab: "x \<le> ab - c \<Longrightarrow> c \<le> ab \<Longrightarrow> x \<le> x + c"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3107
  for x ab c :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3108
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3109
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3110
lemma le_minus': "a + c \<le> b \<Longrightarrow> a \<le> a + c \<Longrightarrow> c \<le> b - a"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3111
  for a b c :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3112
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3113
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3114
lemma le_plus': "a \<le> b \<Longrightarrow> c \<le> b - a \<Longrightarrow> a + c \<le> b"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3115
  for a b c :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3116
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3117
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3118
lemmas le_plus = le_plus' [rotated]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3119
46011
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  3120
lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3121
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3122
lemma word_plus_mono_right: "y \<le> z \<Longrightarrow> x \<le> x + z \<Longrightarrow> x + y \<le> x + z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3123
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3124
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3125
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3126
lemma word_less_minus_cancel: "y - x < z - x \<Longrightarrow> x \<le> z \<Longrightarrow> y < z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3127
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3128
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3129
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3130
lemma word_less_minus_mono_left: "y < z \<Longrightarrow> x \<le> y \<Longrightarrow> y - x < z - x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3131
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3132
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3133
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3134
lemma word_less_minus_mono: "a < c \<Longrightarrow> d < b \<Longrightarrow> a - b < a \<Longrightarrow> c - d < c \<Longrightarrow> a - b < c - d"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3135
  for a b c d :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3136
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3137
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3138
lemma word_le_minus_cancel: "y - x \<le> z - x \<Longrightarrow> x \<le> z \<Longrightarrow> y \<le> z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3139
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3140
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3141
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3142
lemma word_le_minus_mono_left: "y \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> y - x \<le> z - x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3143
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3144
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3145
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3146
lemma word_le_minus_mono:
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3147
  "a \<le> c \<Longrightarrow> d \<le> b \<Longrightarrow> a - b \<le> a \<Longrightarrow> c - d \<le> c \<Longrightarrow> a - b \<le> c - d"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3148
  for a b c d :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3149
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3150
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3151
lemma plus_le_left_cancel_wrap: "x + y' < x \<Longrightarrow> x + y < x \<Longrightarrow> x + y' < x + y \<longleftrightarrow> y' < y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3152
  for x y y' :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3153
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3154
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3155
lemma plus_le_left_cancel_nowrap: "x \<le> x + y' \<Longrightarrow> x \<le> x + y \<Longrightarrow> x + y' < x + y \<longleftrightarrow> y' < y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3156
  for x y y' :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3157
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3158
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3159
lemma word_plus_mono_right2: "a \<le> a + b \<Longrightarrow> c \<le> b \<Longrightarrow> a \<le> a + c"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3160
  for a b c :: "'a::len word"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3161
  by uint_arith
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3162
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3163
lemma word_less_add_right: "x < y - z \<Longrightarrow> z \<le> y \<Longrightarrow> x + z < y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3164
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3165
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3166
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3167
lemma word_less_sub_right: "x < y + z \<Longrightarrow> y \<le> x \<Longrightarrow> x - y < z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3168
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3169
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3170
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3171
lemma word_le_plus_either: "x \<le> y \<or> x \<le> z \<Longrightarrow> y \<le> y + z \<Longrightarrow> x \<le> y + z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3172
  for x y z :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3173
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3174
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3175
lemma word_less_nowrapI: "x < z - k \<Longrightarrow> k \<le> z \<Longrightarrow> 0 < k \<Longrightarrow> x < x + k"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3176
  for x z k :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3177
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3178
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3179
lemma inc_le: "i < m \<Longrightarrow> i + 1 \<le> m"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3180
  for i m :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3181
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3182
82687
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3183
lemma less_imp_less_eq_dec:
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3184
  \<open>v \<le> w - 1\<close> if \<open>v < w\<close> for v w :: \<open>'a::len word\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3185
using that proof transfer
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3186
  show \<open>take_bit LENGTH('a) k \<le> take_bit LENGTH('a) (l - 1)\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3187
    if \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3188
    for k l :: int
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3189
    using that by (cases \<open>take_bit LENGTH('a) l = 0\<close>)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3190
      (auto simp add: take_bit_decr_eq)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3191
qed
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3192
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3193
lemma inc_less_eq_triv_imp:
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3194
  \<open>w = - 1\<close> if \<open>w + 1 \<le> w\<close> for w :: \<open>'a::len word\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3195
proof (rule ccontr)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3196
  assume \<open>w \<noteq> - 1\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3197
  with that show False
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3198
    by transfer (auto simp add: take_bit_eq_mask_iff dest: take_bit_decr_eq)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3199
qed
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3200
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3201
lemma less_eq_dec_triv_imp:
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3202
  \<open>w = 0\<close> if \<open>w \<le> w - 1\<close> for w :: \<open>'a::len word\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3203
proof (rule ccontr)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3204
  assume \<open>w \<noteq> 0\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3205
  with that show False
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3206
    by transfer (auto simp add: take_bit_eq_mask_iff dest: take_bit_decr_eq)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3207
qed
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3208
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3209
lemma inc_less_eq_iff:
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3210
  \<open>v + 1 \<le> w \<longleftrightarrow> v = - 1 \<or> v < w\<close> for v w :: \<open>'a::len word\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3211
  by (auto intro: inc_less_eq_triv_imp inc_le)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3212
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3213
lemma less_eq_dec_iff:
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3214
  \<open>v \<le> w - 1 \<longleftrightarrow> w = 0 \<or> v < w\<close> for v w :: \<open>'a::len word\<close>
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3215
  by (auto intro: less_eq_dec_triv_imp less_imp_less_eq_dec)
97b9ac57751e some more lemmas
haftmann
parents: 82526
diff changeset
  3216
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3217
lemma inc_i: "1 \<le> i \<Longrightarrow> i < m \<Longrightarrow> 1 \<le> i + 1 \<and> i + 1 \<le> m"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3218
  for i m :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3219
  by uint_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3220
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3221
lemma udvd_incr_lem:
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3222
  "\<lbrakk>up < uq; up = ua + n * uint K; uq = ua + n' * uint K\<rbrakk>
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3223
    \<Longrightarrow> up + uint K \<le> uq"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3224
  by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3225
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3226
lemma udvd_incr':
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3227
  "p < q \<Longrightarrow> uint p = ua + n * uint K \<Longrightarrow>
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3228
    uint q = ua + n' * uint K \<Longrightarrow> p + K \<le> q"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3229
  unfolding word_less_alt word_le_def
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3230
  by (metis (full_types) order_trans udvd_incr_lem uint_add_le)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3231
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3232
lemma udvd_decr':
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3233
  assumes "p < q" "uint p = ua + n * uint K" "uint q = ua + n' * uint K"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3234
    shows "uint q = ua + n' * uint K \<Longrightarrow> p \<le> q - K"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3235
proof -
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3236
  have "\<And>w v. uint (w::'a word) \<le> uint v + uint (w - v)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3237
    by (metis (no_types) add_diff_cancel_left' diff_add_cancel uint_add_le)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3238
  moreover have "uint K + uint p \<le> uint q"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3239
    using assms by (metis (no_types) add_diff_cancel_left' diff_add_cancel udvd_incr_lem word_less_def)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3240
  ultimately show ?thesis
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3241
    by (meson add_le_cancel_left order_trans word_less_eq_iff_unsigned)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3242
qed
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3243
45816
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  3244
lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left]
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  3245
lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left]
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  3246
lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3247
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3248
lemma udvd_minus_le': "xy < k \<Longrightarrow> z udvd xy \<Longrightarrow> z udvd k \<Longrightarrow> xy \<le> k - z"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3249
  unfolding udvd_unfold_int
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3250
  by (meson udvd_decr0)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3251
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3252
lemma udvd_incr2_K:
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3253
  "p < a + s \<Longrightarrow> a \<le> a + s \<Longrightarrow> K udvd s \<Longrightarrow> K udvd p - a \<Longrightarrow> a \<le> p \<Longrightarrow>
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3254
    0 < K \<Longrightarrow> p \<le> p + K \<and> p + K \<le> a + s"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3255
  unfolding udvd_unfold_int
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3256
  by (smt (verit, best) diff_add_cancel leD udvd_incr_lem uint_plus_if'
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3257
      word_less_eq_iff_unsigned word_sub_le)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3258
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3259
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3260
subsection \<open>Arithmetic type class instantiations\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3261
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3262
lemmas word_le_0_iff [simp] =
70749
5d06b7bb9d22 More type class generalisations. Note that linorder_antisym_conv1 and linorder_antisym_conv2 no longer exist.
paulson <lp15@cam.ac.uk>
parents: 70342
diff changeset
  3263
  word_zero_le [THEN leD, THEN antisym_conv1]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3264
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3265
lemma word_of_int_nat: "0 \<le> x \<Longrightarrow> word_of_int x = of_nat (nat x)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3266
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3267
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  3268
text \<open>
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  3269
  note that \<open>iszero_def\<close> is only for class \<open>comm_semiring_1_cancel\<close>,
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  3270
  which requires word length \<open>\<ge> 1\<close>, ie \<open>'a::len word\<close>
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  3271
\<close>
46603
83a5160e6b4d removed unnecessary lemma zero_bintrunc
huffman
parents: 46602
diff changeset
  3272
lemma iszero_word_no [simp]:
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3273
  "iszero (numeral bin :: 'a::len word) =
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3274
    iszero (take_bit LENGTH('a) (numeral bin :: int))"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3275
  by (metis iszero_def uint_0_iff uint_bintrunc)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3276
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3277
text \<open>Use \<open>iszero\<close> to simplify equalities between word numerals.\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3278
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3279
lemmas word_eq_numeral_iff_iszero [simp] =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3280
  eq_numeral_iff_iszero [where 'a="'a::len word"]
46603
83a5160e6b4d removed unnecessary lemma zero_bintrunc
huffman
parents: 46602
diff changeset
  3281
79590
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3282
lemma word_less_eq_imp_half_less_eq:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3283
  \<open>v div 2 \<le> w div 2\<close> if \<open>v \<le> w\<close> for v w :: \<open>'a::len word\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3284
  using that by (simp add: word_le_nat_alt unat_div div_le_mono)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3285
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3286
lemma word_half_less_imp_less_eq:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3287
  \<open>v \<le> w\<close> if \<open>v div 2 < w div 2\<close> for v w :: \<open>'a::len word\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3288
  using that linorder_linear word_less_eq_imp_half_less_eq by fastforce
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  3289
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3290
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3291
subsection \<open>Word and nat\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3292
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3293
lemma word_nchotomy: "\<forall>w :: 'a::len word. \<exists>n. w = of_nat n \<and> n < 2 ^ LENGTH('a)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3294
  by (metis of_nat_unat ucast_id unsigned_less)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3295
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3296
lemma of_nat_eq: "of_nat n = w \<longleftrightarrow> (\<exists>q. n = unat w + q * 2 ^ LENGTH('a))"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3297
  for w :: "'a::len word"
68157
057d5b4ce47e removed some non-essential rules
haftmann
parents: 67443
diff changeset
  3298
  using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric]
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
  3299
  by (auto simp flip: take_bit_eq_mod simp add: unsigned_of_nat)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3300
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3301
lemma of_nat_eq_size: "of_nat n = w \<longleftrightarrow> (\<exists>q. n = unat w + q * 2 ^ size w)"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3302
  unfolding word_size by (rule of_nat_eq)
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3303
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3304
lemma of_nat_0: "of_nat m = (0::'a::len word) \<longleftrightarrow> (\<exists>q. m = q * 2 ^ LENGTH('a))"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3305
  by (simp add: of_nat_eq)
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3306
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3307
lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)"
45805
3c609e8785f2 tidied Word.thy;
huffman
parents: 45804
diff changeset
  3308
  by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]])
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3309
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3310
lemma of_nat_gt_0: "of_nat k \<noteq> 0 \<Longrightarrow> 0 < k"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3311
  by (cases k) auto
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3312
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3313
lemma of_nat_neq_0: "0 < k \<Longrightarrow> k < 2 ^ LENGTH('a::len) \<Longrightarrow> of_nat k \<noteq> (0 :: 'a word)"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3314
  by (auto simp : of_nat_0)
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3315
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3316
lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3317
  by simp
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3318
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3319
lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3320
  by (simp add: wi_hom_mult)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3321
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3322
lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3323
  by transfer (simp add: ac_simps)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3324
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3325
lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0"
45995
b16070689726 declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents: 45958
diff changeset
  3326
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3327
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3328
lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)"
45995
b16070689726 declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents: 45958
diff changeset
  3329
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3330
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3331
lemmas Abs_fnat_homs =
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3332
  Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3333
  Abs_fnat_hom_0 Abs_fnat_hom_1
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3334
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3335
lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3336
  by simp
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3337
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3338
lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3339
  by simp
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3340
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3341
lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3342
  by (subst Abs_fnat_hom_Suc [symmetric]) simp
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3343
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3344
lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3345
  by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3346
  
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3347
lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3348
  by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3349
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3350
lemmas word_arith_nat_defs =
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3351
  word_arith_nat_add word_arith_nat_mult
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3352
  word_arith_nat_Suc Abs_fnat_hom_0
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3353
  Abs_fnat_hom_1 word_arith_nat_div
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3354
  word_arith_nat_mod
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3355
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3356
lemma unat_of_nat:
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3357
  \<open>unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3358
  by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3359
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3360
lemma unat_cong: "x = y \<Longrightarrow> unat x = unat y"
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3361
  by (fact arg_cong)
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3362
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3363
lemmas unat_word_ariths = word_arith_nat_defs
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3364
  [THEN trans [OF unat_cong unat_of_nat]]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3365
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3366
lemmas word_sub_less_iff = word_sub_le_iff
45816
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  3367
  [unfolded linorder_not_less [symmetric] Not_eq_iff]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3368
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3369
lemma unat_add_lem:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3370
  "unat x + unat y < 2 ^ LENGTH('a) \<longleftrightarrow> unat (x + y) = unat x + unat y"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3371
  for x y :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3372
  by (metis mod_less unat_word_ariths(1) unsigned_less)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3373
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3374
lemma unat_mult_lem:
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3375
  "unat x * unat y < 2 ^ LENGTH('a) \<longleftrightarrow> unat (x * y) = unat x * unat y"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  3376
  for x y :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3377
  by (metis mod_less unat_word_ariths(2) unsigned_less)
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3378
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3379
lemma le_no_overflow: "x \<le> b \<Longrightarrow> a \<le> a + b \<Longrightarrow> x \<le> a + b"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3380
  for a b x :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3381
  using word_le_plus_either by blast
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3382
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3383
lemma uint_div:
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3384
  \<open>uint (x div y) = uint x div uint y\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3385
  by (fact uint_div_distrib)
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3386
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3387
lemma uint_mod:
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3388
  \<open>uint (x mod y) = uint x mod uint y\<close>
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3389
  by (fact uint_mod_distrib)
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3390
  
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3391
lemma no_plus_overflow_unat_size: "x \<le> x + y \<longleftrightarrow> unat x + unat y < 2 ^ size x"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3392
  for x y :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3393
  unfolding word_size by unat_arith
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3394
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3395
lemmas no_olen_add_nat =
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3396
  no_plus_overflow_unat_size [unfolded word_size]
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3397
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3398
lemmas unat_plus_simple =
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3399
  trans [OF no_olen_add_nat unat_add_lem]
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3400
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3401
lemma word_div_mult: "\<lbrakk>0 < y; unat x * unat y < 2 ^ LENGTH('a)\<rbrakk> \<Longrightarrow> x * y div y = x"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3402
  for x y :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3403
  by (simp add: unat_eq_zero unat_mult_lem word_arith_nat_div)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3404
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3405
lemma div_lt': "i \<le> k div x \<Longrightarrow> unat i * unat x < 2 ^ LENGTH('a)"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3406
  for i k x :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3407
  by unat_arith (meson le_less_trans less_mult_imp_div_less not_le unsigned_less)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3408
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3409
lemmas div_lt'' = order_less_imp_le [THEN div_lt']
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3410
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3411
lemma div_lt_mult: "\<lbrakk>i < k div x; 0 < x\<rbrakk> \<Longrightarrow> i * x < k"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3412
  for i k x :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3413
  by (metis div_le_mono div_lt'' not_le unat_div word_div_mult word_less_iff_unsigned)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3414
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3415
lemma div_le_mult: "\<lbrakk>i \<le> k div x; 0 < x\<rbrakk> \<Longrightarrow> i * x \<le> k"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3416
  for i k x :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3417
  by (metis div_lt' less_mult_imp_div_less not_less unat_arith_simps(2) unat_div unat_mult_lem)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3418
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3419
lemma div_lt_uint': "i \<le> k div x \<Longrightarrow> uint i * uint x < 2 ^ LENGTH('a)"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3420
  for i k x :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3421
  unfolding uint_nat
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3422
  by (metis div_lt' int_ops(7) of_nat_unat uint_mult_lem unat_mult_lem)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3423
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3424
lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint']
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3425
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  3426
lemma word_le_exists': "x \<le> y \<Longrightarrow> \<exists>z. y = x + z \<and> uint x + uint z < 2 ^ LENGTH('a)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3427
  for x y z :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3428
  by (metis add.commute diff_add_cancel no_olen_add)
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3429
  
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3430
lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3431
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3432
lemmas plus_minus_no_overflow =
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3433
  order_less_imp_le [THEN plus_minus_no_overflow_ab]
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3434
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3435
lemmas mcs = word_less_minus_cancel word_less_minus_mono_left
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3436
  word_le_minus_cancel word_le_minus_mono_left
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3437
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3438
lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3439
lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3440
lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3441
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3442
lemma le_unat_uoi:
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3443
  \<open>y \<le> unat z \<Longrightarrow> unat (word_of_nat y :: 'a word) = y\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3444
  for z :: \<open>'a::len word\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  3445
  by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3446
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66453
diff changeset
  3447
lemmas thd = times_div_less_eq_dividend
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3448
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3449
lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3450
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3451
lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3452
  for n b :: "'a::len word"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3453
  by (fact div_mult_mod_eq)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3454
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3455
lemma word_div_mult_le: "a div b * b \<le> a"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3456
  for a b :: "'a::len word"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3457
  by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3458
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3459
lemma word_mod_less_divisor: "0 < n \<Longrightarrow> m mod n < n"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3460
  for m n :: "'a::len word"
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3461
  by (simp add: unat_arith_simps)
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3462
  
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3463
lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)"
45995
b16070689726 declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents: 45958
diff changeset
  3464
  by (induct n) (simp_all add: wi_hom_mult [symmetric])
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3465
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3466
lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3467
  by (simp add : word_of_int_power_hom [symmetric])
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3468
70183
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3469
lemma unatSuc: "1 + n \<noteq> 0 \<Longrightarrow> unat (1 + n) = Suc (unat n)"
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3470
  for n :: "'a::len word"
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3471
  by unat_arith
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3472
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3473
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3474
subsection \<open>Cardinality, finiteness of set of words\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3475
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3476
lemma inj_on_word_of_int: \<open>inj_on (word_of_int :: int \<Rightarrow> 'a word) {0..<2 ^ LENGTH('a::len)}\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3477
  unfolding inj_on_def
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3478
  by (metis atLeastLessThan_iff word_of_int_inverse)
71948
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
  3479
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3480
lemma range_uint: \<open>range (uint :: 'a word \<Rightarrow> int) = {0..<2 ^ LENGTH('a::len)}\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  3481
  apply transfer
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3482
  apply (auto simp: image_iff)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  3483
  apply (metis take_bit_int_eq_self_iff)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  3484
  done
71948
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
  3485
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3486
lemma UNIV_eq: \<open>(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3487
  by (auto simp: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split)
45809
2bee94cbae72 finite class instance for word type; remove unused lemmas
huffman
parents: 45808
diff changeset
  3488
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3489
lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)"
71948
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
  3490
  by (simp add: UNIV_eq card_image inj_on_word_of_int)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3491
70183
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3492
lemma card_word_size: "CARD('a word) = 2 ^ size x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3493
  for x :: "'a::len word"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3494
  unfolding word_size by (rule card_word)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3495
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  3496
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  3497
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3498
instance word :: (len) finite
71948
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
  3499
  by standard (simp add: UNIV_eq)
6ede899d26d3 fundamental construction of word type following existing transfer rules
haftmann
parents: 71947
diff changeset
  3500
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3501
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3502
subsection \<open>Bitwise Operations on Words\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3503
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  3504
context
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  3505
  includes bit_operations_syntax
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  3506
begin
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  3507
46011
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  3508
lemma word_wi_log_defs:
71149
a7d1fb0c9e16 proper prefix syntax
haftmann
parents: 70901
diff changeset
  3509
  "NOT (word_of_int a) = word_of_int (NOT a)"
46011
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  3510
  "word_of_int a AND word_of_int b = word_of_int (a AND b)"
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  3511
  "word_of_int a OR word_of_int b = word_of_int (a OR b)"
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  3512
  "word_of_int a XOR word_of_int b = word_of_int (a XOR b)"
47374
9475d524bafb set up and use lift_definition for word operations
huffman
parents: 47372
diff changeset
  3513
  by (transfer, rule refl)+
47372
9ab4e22dac7b configure transfer method for 'a word -> int
huffman
parents: 47168
diff changeset
  3514
46011
96a5f44c22da replace 'lemmas' with explicit 'lemma'
huffman
parents: 46010
diff changeset
  3515
lemma word_no_log_defs [simp]:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3516
  "NOT (numeral a) = word_of_int (NOT (numeral a))"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3517
  "NOT (- numeral a) = word_of_int (NOT (- numeral a))"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3518
  "numeral a AND numeral b = word_of_int (numeral a AND numeral b)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3519
  "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3520
  "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3521
  "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3522
  "numeral a OR numeral b = word_of_int (numeral a OR numeral b)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3523
  "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3524
  "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3525
  "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3526
  "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3527
  "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3528
  "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3529
  "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)"
47372
9ab4e22dac7b configure transfer method for 'a word -> int
huffman
parents: 47168
diff changeset
  3530
  by (transfer, rule refl)+
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3531
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3532
text \<open>Special cases for when one of the arguments equals 1.\<close>
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46057
diff changeset
  3533
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46057
diff changeset
  3534
lemma word_bitwise_1_simps [simp]:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3535
  "NOT (1::'a::len word) = -2"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3536
  "1 AND numeral b = word_of_int (1 AND numeral b)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3537
  "1 AND - numeral b = word_of_int (1 AND - numeral b)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3538
  "numeral a AND 1 = word_of_int (numeral a AND 1)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3539
  "- numeral a AND 1 = word_of_int (- numeral a AND 1)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3540
  "1 OR numeral b = word_of_int (1 OR numeral b)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3541
  "1 OR - numeral b = word_of_int (1 OR - numeral b)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3542
  "numeral a OR 1 = word_of_int (numeral a OR 1)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3543
  "- numeral a OR 1 = word_of_int (- numeral a OR 1)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3544
  "1 XOR numeral b = word_of_int (1 XOR numeral b)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3545
  "1 XOR - numeral b = word_of_int (1 XOR - numeral b)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3546
  "numeral a XOR 1 = word_of_int (numeral a XOR 1)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54225
diff changeset
  3547
  "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)"
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
  3548
              apply (simp_all add: word_uint_eq_iff unsigned_not_eq unsigned_and_eq unsigned_or_eq
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
  3549
         unsigned_xor_eq of_nat_take_bit ac_simps unsigned_of_int)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3550
       apply (simp_all add: minus_numeral_eq_not_sub_one)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3551
   apply (simp_all only: sub_one_eq_not_neg bit.xor_compl_right take_bit_xor bit.double_compl)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3552
   apply simp_all
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3553
  done
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46057
diff changeset
  3554
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3555
text \<open>Special cases for when one of the arguments equals -1.\<close>
56979
376604d56b54 added lemmas for -1
noschinl
parents: 56078
diff changeset
  3556
376604d56b54 added lemmas for -1
noschinl
parents: 56078
diff changeset
  3557
lemma word_bitwise_m1_simps [simp]:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3558
  "NOT (-1::'a::len word) = 0"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3559
  "(-1::'a::len word) AND x = x"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3560
  "x AND (-1::'a::len word) = x"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3561
  "(-1::'a::len word) OR x = -1"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3562
  "x OR (-1::'a::len word) = -1"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3563
  " (-1::'a::len word) XOR x = NOT x"
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3564
  "x XOR (-1::'a::len word) = NOT x"
56979
376604d56b54 added lemmas for -1
noschinl
parents: 56078
diff changeset
  3565
  by (transfer, simp)+
376604d56b54 added lemmas for -1
noschinl
parents: 56078
diff changeset
  3566
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3567
lemma word_of_int_not_numeral_eq [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3568
  \<open>(word_of_int (NOT (numeral bin)) :: 'a::len word) = - numeral bin - 1\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3569
  by transfer (simp add: not_eq_complement)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74108
diff changeset
  3570
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3571
lemma uint_and:
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3572
  \<open>uint (x AND y) = uint x AND uint y\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3573
  by transfer simp
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3574
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3575
lemma uint_or:
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3576
  \<open>uint (x OR y) = uint x OR uint y\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3577
  by transfer simp
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3578
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3579
lemma uint_xor:
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3580
  \<open>uint (x XOR y) = uint x XOR uint y\<close>
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3581
  by transfer simp
47372
9ab4e22dac7b configure transfer method for 'a word -> int
huffman
parents: 47168
diff changeset
  3582
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  3583
\<comment> \<open>get from commutativity, associativity etc of \<open>int_and\<close> etc to same for \<open>word_and etc\<close>\<close>
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3584
lemmas bwsimps =
46013
d2f179d26133 remove some duplicate lemmas
huffman
parents: 46012
diff changeset
  3585
  wi_hom_add
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3586
  word_wi_log_defs
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3587
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3588
lemma word_bw_assocs:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3589
  "(x AND y) AND z = x AND y AND z"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3590
  "(x OR y) OR z = x OR y OR z"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3591
  "(x XOR y) XOR z = x XOR y XOR z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3592
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3593
  by (fact ac_simps)+
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3594
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3595
lemma word_bw_comms:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3596
  "x AND y = y AND x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3597
  "x OR y = y OR x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3598
  "x XOR y = y XOR x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3599
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3600
  by (fact ac_simps)+
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3601
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3602
lemma word_bw_lcs:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3603
  "y AND x AND z = x AND y AND z"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3604
  "y OR x OR z = x OR y OR z"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3605
  "y XOR x XOR z = x XOR y XOR z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3606
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3607
  by (fact ac_simps)+
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3608
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3609
lemma word_log_esimps:
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3610
  "x AND 0 = 0"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3611
  "x AND -1 = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3612
  "x OR 0 = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3613
  "x OR -1 = -1"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3614
  "x XOR 0 = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3615
  "x XOR -1 = NOT x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3616
  "0 AND x = 0"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3617
  "-1 AND x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3618
  "0 OR x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3619
  "-1 OR x = -1"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3620
  "0 XOR x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3621
  "-1 XOR x = NOT x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3622
  for x :: "'a::len word"
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3623
  by simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3624
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3625
lemma word_not_dist:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3626
  "NOT (x OR y) = NOT x AND NOT y"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3627
  "NOT (x AND y) = NOT x OR NOT y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3628
  for x :: "'a::len word"
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3629
  by simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3630
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3631
lemma word_bw_same:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3632
  "x AND x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3633
  "x OR x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3634
  "x XOR x = 0"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3635
  for x :: "'a::len word"
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3636
  by simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3637
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3638
lemma word_ao_absorbs [simp]:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3639
  "x AND (y OR x) = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3640
  "x OR y AND x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3641
  "x AND (x OR y) = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3642
  "y AND x OR x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3643
  "(y OR x) AND x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3644
  "x OR x AND y = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3645
  "(x OR y) AND x = x"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3646
  "x AND y OR x = x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3647
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3648
  by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3649
71149
a7d1fb0c9e16 proper prefix syntax
haftmann
parents: 70901
diff changeset
  3650
lemma word_not_not [simp]: "NOT (NOT x) = x"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3651
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3652
  by (fact bit.double_compl)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3653
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3654
lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3655
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3656
  by (fact bit.conj_disj_distrib2)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3657
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3658
lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3659
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3660
  by (fact bit.disj_conj_distrib2)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3661
  
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3662
lemma word_add_not [simp]: "x + NOT x = -1"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3663
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3664
  by (simp add: not_eq_complement)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  3665
  
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3666
lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3667
  for x :: "'a::len word"
47372
9ab4e22dac7b configure transfer method for 'a word -> int
huffman
parents: 47168
diff changeset
  3668
  by transfer (simp add: plus_and_or)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3669
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3670
lemma leoa: "w = x OR y \<Longrightarrow> y = w AND y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3671
  for x :: "'a::len word"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3672
  by auto
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3673
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3674
lemma leao: "w' = x' AND y' \<Longrightarrow> x' = x' OR w'"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3675
  for x' :: "'a::len word"
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3676
  by auto
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3677
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3678
lemma word_ao_equiv: "w = w OR w' \<longleftrightarrow> w' = w AND w'"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3679
  for w w' :: "'a::len word"
48196
b7313810b6e6 explicit is better than implicit;
wenzelm
parents: 47941
diff changeset
  3680
  by (auto intro: leoa leao)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3681
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3682
lemma le_word_or2: "x \<le> x OR y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  3683
  for x y :: "'a::len word"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  3684
  by (simp add: or_greater_eq uint_or word_le_def)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3685
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3686
lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2]
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3687
lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2]
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3688
lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3689
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  3690
lemma bit_horner_sum_bit_word_iff [bit_simps]:
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3691
  \<open>bit (horner_sum of_bool (2 :: 'a::len word) bs) n
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3692
    \<longleftrightarrow> n < min LENGTH('a) (length bs) \<and> bs ! n\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3693
  by transfer (simp add: bit_horner_sum_bit_iff)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3694
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3695
definition word_reverse :: \<open>'a::len word \<Rightarrow> 'a word\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3696
  where \<open>word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0..<LENGTH('a)]))\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3697
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  3698
lemma bit_word_reverse_iff [bit_simps]:
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  3699
  \<open>bit (word_reverse w) n \<longleftrightarrow> n < LENGTH('a) \<and> bit w (LENGTH('a) - Suc n)\<close>
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  3700
  for w :: \<open>'a::len word\<close>
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3701
  by (cases \<open>n < LENGTH('a)\<close>)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3702
    (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3703
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3704
lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3705
  by (rule bit_word_eqI)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3706
    (auto simp: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc)
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3707
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3708
lemma word_rev_gal: "word_reverse w = u \<Longrightarrow> word_reverse u = w"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3709
  by (metis word_rev_rev)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3710
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3711
lemma word_rev_gal': "u = word_reverse w \<Longrightarrow> w = word_reverse u"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3712
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3713
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  3714
lemma word_eq_reverseI:
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  3715
  \<open>v = w\<close> if \<open>word_reverse v = word_reverse w\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3716
  by (metis that word_rev_rev)
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  3717
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3718
lemma uint_2p: "(0::'a::len word) < 2 ^ n \<Longrightarrow> uint (2 ^ n::'a::len word) = 2 ^ n"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3719
  by (cases \<open>n < LENGTH('a)\<close>; transfer; force)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3720
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  3721
lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3722
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3723
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3724
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3725
subsubsection \<open>shift functions in terms of lists of bools\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3726
71997
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3727
lemma drop_bit_word_numeral [simp]:
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3728
  \<open>drop_bit (numeral n) (numeral k) =
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3729
    (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3730
  by transfer simp
4a013c92a091 factored out auxiliary theory
haftmann
parents: 71996
diff changeset
  3731
74498
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3732
lemma drop_bit_word_Suc_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3733
  \<open>drop_bit (Suc n) (numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3734
    (word_of_int (drop_bit (Suc n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3735
  by transfer simp
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3736
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3737
lemma drop_bit_word_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3738
  \<open>drop_bit (numeral n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3739
    (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3740
  by transfer simp
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3741
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3742
lemma drop_bit_word_Suc_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3743
  \<open>drop_bit (Suc n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3744
    (word_of_int (drop_bit (Suc n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3745
  by transfer simp
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3746
73853
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  3747
lemma signed_drop_bit_word_numeral [simp]:
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  3748
  \<open>signed_drop_bit (numeral n) (numeral k) =
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  3749
    (word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k))) :: 'a::len word)\<close>
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  3750
  by transfer simp
52b829b18066 more lemmas
haftmann
parents: 73816
diff changeset
  3751
74498
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3752
lemma signed_drop_bit_word_Suc_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3753
  \<open>signed_drop_bit (Suc n) (numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3754
    (word_of_int (drop_bit (Suc n) (signed_take_bit (LENGTH('a) - 1) (numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3755
  by transfer simp
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3756
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3757
lemma signed_drop_bit_word_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3758
  \<open>signed_drop_bit (numeral n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3759
    (word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (- numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3760
  by transfer simp
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3761
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3762
lemma signed_drop_bit_word_Suc_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3763
  \<open>signed_drop_bit (Suc n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3764
    (word_of_int (drop_bit (Suc n) (signed_take_bit (LENGTH('a) - 1) (- numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3765
  by transfer simp
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3766
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3767
lemma take_bit_word_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3768
  \<open>take_bit (numeral n) (numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3769
    (word_of_int (take_bit (min LENGTH('a) (numeral n)) (numeral k)) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3770
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3771
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3772
lemma take_bit_word_Suc_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3773
  \<open>take_bit (Suc n) (numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3774
    (word_of_int (take_bit (min LENGTH('a) (Suc n)) (numeral k)) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3775
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3776
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3777
lemma take_bit_word_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3778
  \<open>take_bit (numeral n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3779
    (word_of_int (take_bit (min LENGTH('a) (numeral n)) (- numeral k)) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3780
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3781
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3782
lemma take_bit_word_Suc_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3783
  \<open>take_bit (Suc n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3784
    (word_of_int (take_bit (min LENGTH('a) (Suc n)) (- numeral k)) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3785
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3786
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3787
lemma signed_take_bit_word_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3788
  \<open>signed_take_bit (numeral n) (numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3789
    (word_of_int (signed_take_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3790
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3791
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3792
lemma signed_take_bit_word_Suc_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3793
  \<open>signed_take_bit (Suc n) (numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3794
    (word_of_int (signed_take_bit (Suc n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3795
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3796
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3797
lemma signed_take_bit_word_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3798
  \<open>signed_take_bit (numeral n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3799
    (word_of_int (signed_take_bit (numeral n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3800
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3801
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3802
lemma signed_take_bit_word_Suc_minus_numeral [simp]:
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3803
  \<open>signed_take_bit (Suc n) (- numeral k) =
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3804
    (word_of_int (signed_take_bit (Suc n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3805
  by transfer rule
27475e64a887 more complete simp rules
haftmann
parents: 74496
diff changeset
  3806
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3807
lemma False_map2_or: "\<lbrakk>set xs \<subseteq> {False}; length ys = length xs\<rbrakk> \<Longrightarrow> map2 (\<or>) xs ys = ys"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3808
  by (induction xs arbitrary: ys) (auto simp: length_Suc_conv)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3809
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3810
lemma align_lem_or:
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3811
  assumes "length xs = n + m" "length ys = n + m" 
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3812
    and "drop m xs = replicate n False" "take m ys = replicate m False"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3813
  shows "map2 (\<or>) xs ys = take m xs @ drop m ys"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3814
  using assms
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3815
proof (induction xs arbitrary: ys m)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3816
  case (Cons a xs)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3817
  then show ?case
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3818
    by (cases m) (auto simp: length_Suc_conv False_map2_or)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3819
qed auto
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3820
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3821
lemma False_map2_and: "\<lbrakk>set xs \<subseteq> {False}; length ys = length xs\<rbrakk> \<Longrightarrow> map2 (\<and>) xs ys = xs"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3822
  by (induction xs arbitrary: ys) (auto simp: length_Suc_conv)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3823
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3824
lemma align_lem_and:
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3825
  assumes "length xs = n + m" "length ys = n + m" 
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3826
    and "drop m xs = replicate n False" "take m ys = replicate m False"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3827
  shows "map2 (\<and>) xs ys = replicate (n + m) False"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3828
  using assms
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3829
proof (induction xs arbitrary: ys m)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3830
  case (Cons a xs)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3831
  then show ?case
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3832
    by (cases m) (auto simp: length_Suc_conv set_replicate_conv_if False_map2_and)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3833
qed auto
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3834
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  3835
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  3836
subsubsection \<open>Mask\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3837
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  3838
lemma minus_1_eq_mask:
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3839
  \<open>- 1 = (mask LENGTH('a) :: 'a::len word)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3840
  by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff)
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  3841
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  3842
lemma mask_eq_decr_exp:
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3843
  \<open>mask n = 2 ^ n - (1 :: 'a::len word)\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3844
  by (fact mask_eq_exp_minus_1)
71953
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3845
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3846
lemma mask_Suc_rec:
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3847
  \<open>mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3848
  by (simp add: mask_eq_exp_minus_1)
71953
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3849
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3850
context
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3851
begin
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3852
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  3853
qualified lemma bit_mask_iff [bit_simps]:
71990
66beb9d92e43 explicit proofs for bit projections
haftmann
parents: 71986
diff changeset
  3854
  \<open>bit (mask m :: 'a::len word) n \<longleftrightarrow> n < min LENGTH('a) m\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3855
  by (simp add: bit_mask_iff not_le)
71953
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3856
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3857
end
428609096812 more lemmas and less name space pollution
haftmann
parents: 71952
diff changeset
  3858
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3859
lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3860
  by transfer simp 
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3861
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3862
lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  3863
  by transfer (simp add: ac_simps take_bit_eq_mask)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3864
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3865
lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)"
74496
807b094a9b78 avoid overaggressive contraction of conversions
haftmann
parents: 74391
diff changeset
  3866
  by (simp add: take_bit_eq_mask of_int_and_eq of_int_mask_eq)
46023
fad87bb608fc restate some lemmas to respect int/bin distinction
huffman
parents: 46022
diff changeset
  3867
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3868
lemma and_mask_wi':
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3869
  "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3870
  by (auto simp: and_mask_wi min_def wi_bintr)
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 64243
diff changeset
  3871
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3872
lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  3873
  unfolding word_numeral_alt by (rule and_mask_wi)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3874
45811
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  3875
lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)"
72128
3707cf7b370b reduced prominence od theory Bits_Int
haftmann
parents: 72102
diff changeset
  3876
  by (simp only: and_mask_bintr take_bit_eq_mod)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3877
72130
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  3878
lemma uint_mask_eq:
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  3879
  \<open>uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\<close>
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  3880
  by transfer simp
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  3881
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3882
lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3883
  by (metis take_bit_eq_mask take_bit_int_less_exp unsigned_take_bit_eq)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3884
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  3885
lemma mask_eq_iff: "w AND mask n = w \<longleftrightarrow> uint w < 2 ^ n"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3886
  by (metis and_mask_bintr and_mask_lt_2p take_bit_int_eq_self take_bit_nonnegative
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  3887
      uint_sint word_of_int_uint)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3888
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3889
lemma and_mask_dvd: "2 ^ n dvd uint w \<longleftrightarrow> w AND mask n = 0"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3890
  by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3891
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3892
lemma and_mask_dvd_nat: "2 ^ n dvd unat w \<longleftrightarrow> w AND mask n = 0"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3893
  by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3894
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3895
lemma word_2p_lem: "n < size w \<Longrightarrow> w < 2 ^ n = (uint w < 2 ^ n)"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3896
  for w :: "'a::len word"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3897
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3898
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3899
lemma less_mask_eq:
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3900
  fixes x :: "'a::len word"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3901
  assumes "x < 2 ^ n" shows "x AND mask n = x"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3902
  by (metis (no_types) assms lt2p_lem mask_eq_iff not_less word_2p_lem word_size)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3903
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3904
lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]]
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3905
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  3906
lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3907
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3908
lemma and_mask_less_size: "n < size x \<Longrightarrow> x AND mask n < 2 ^ n"
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3909
  for x :: \<open>'a::len word\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3910
  unfolding word_size by (erule and_mask_less')
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3911
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3912
lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \<Longrightarrow> c > 0 \<Longrightarrow> x mod c = x AND mask n"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3913
  for c x :: "'a::len word"
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3914
  by (auto simp: word_mod_def uint_2p and_mask_mod_2p)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3915
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3916
lemma mask_eqs:
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3917
  "(a AND mask n) + b AND mask n = a + b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3918
  "a + (b AND mask n) AND mask n = a + b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3919
  "(a AND mask n) - b AND mask n = a - b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3920
  "a - (b AND mask n) AND mask n = a - b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3921
  "a * (b AND mask n) AND mask n = a * b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3922
  "(b AND mask n) * a AND mask n = b * a AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3923
  "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3924
  "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3925
  "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3926
  "- (a AND mask n) AND mask n = - a AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3927
  "word_succ (a AND mask n) AND mask n = word_succ a AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3928
  "word_pred (a AND mask n) AND mask n = word_pred a AND mask n"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3929
  using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b]
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3930
  unfolding take_bit_eq_mask [symmetric]
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3931
  by (transfer; simp add: take_bit_eq_mod mod_simps)+
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3932
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  3933
lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n"
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  3934
  for x :: \<open>'a::len word\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3935
  using word_of_int_Ex [where x=x]
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3936
  unfolding take_bit_eq_mask [symmetric]
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3937
  by (transfer; simp add: take_bit_eq_mod mod_simps)+
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3938
70183
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3939
lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3940
  by transfer simp
70183
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 70175
diff changeset
  3941
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  3942
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3943
subsubsection \<open>Slices\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3944
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3945
definition slice1 :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'b::len word\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3946
  where \<open>slice1 n w = (if n < LENGTH('a)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3947
    then ucast (drop_bit (LENGTH('a) - n) w)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3948
    else push_bit (n - LENGTH('a)) (ucast w))\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3949
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  3950
lemma bit_slice1_iff [bit_simps]:
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3951
  \<open>bit (slice1 m w :: 'b::len word) n \<longleftrightarrow> m - LENGTH('a) \<le> n \<and> n < min LENGTH('b) m
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3952
    \<and> bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3953
  for w :: \<open>'a::len word\<close>
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  3954
  by (auto simp: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff not_less not_le ac_simps
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3955
    dest: bit_imp_le_length)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3956
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3957
definition slice :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'b::len word\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3958
  where \<open>slice n = slice1 (LENGTH('a) - n)\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3959
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  3960
lemma bit_slice_iff [bit_simps]:
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3961
  \<open>bit (slice m w :: 'b::len word) n \<longleftrightarrow> n < min LENGTH('b) (LENGTH('a) - m) \<and> bit w (n + LENGTH('a) - (LENGTH('a) - m))\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3962
  for w :: \<open>'a::len word\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3963
  by (simp add: slice_def word_size bit_slice1_iff)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3964
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3965
lemma slice1_0 [simp] : "slice1 n 0 = 0"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3966
  unfolding slice1_def by simp
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3967
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3968
lemma slice_0 [simp] : "slice n 0 = 0"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3969
  unfolding slice_def by auto
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3970
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3971
lemma ucast_slice1: "ucast w = slice1 (size w) w"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3972
  unfolding slice1_def by (simp add: size_word.rep_eq)
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3973
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3974
lemma ucast_slice: "ucast w = slice 0 w"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3975
  by (simp add: slice_def slice1_def)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3976
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3977
lemma slice_id: "slice 0 t = t"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3978
  by (simp only: ucast_slice [symmetric] ucast_id)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3979
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3980
lemma rev_slice1:
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3981
  \<open>slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3982
  if \<open>n + k = LENGTH('a) + LENGTH('b)\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3983
proof (rule bit_word_eqI)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3984
  fix m
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3985
  assume *: \<open>m < LENGTH('a)\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3986
  from that have **: \<open>LENGTH('b) = n + k - LENGTH('a)\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3987
    by simp
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3988
  show \<open>bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \<longleftrightarrow> bit (word_reverse (slice1 k w :: 'a word)) m\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3989
    unfolding bit_slice1_iff bit_word_reverse_iff
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3990
    using * **
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3991
    by (cases \<open>n \<le> LENGTH('a)\<close>; cases \<open>k \<le> LENGTH('a)\<close>) auto
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  3992
qed
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3993
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3994
lemma rev_slice:
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3995
  "n + k + LENGTH('a::len) = LENGTH('b::len) \<Longrightarrow>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3996
    slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3997
  unfolding slice_def word_size
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  3998
  by (simp add: rev_slice1)
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  3999
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4000
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4001
subsubsection \<open>Revcast\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4002
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4003
definition revcast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4004
  where \<open>revcast = slice1 LENGTH('b)\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4005
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72515
diff changeset
  4006
lemma bit_revcast_iff [bit_simps]:
72027
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4007
  \<open>bit (revcast w :: 'b::len word) n \<longleftrightarrow> LENGTH('b) - LENGTH('a) \<le> n \<and> n < LENGTH('b)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4008
    \<and> bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4009
  for w :: \<open>'a::len word\<close>
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4010
  by (simp add: revcast_def bit_slice1_iff)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4011
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4012
lemma revcast_slice1 [OF refl]: "rc = revcast w \<Longrightarrow> slice1 (size rc) w = rc"
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4013
  by (simp add: revcast_def word_size)
759532ef0885 prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents: 72010
diff changeset
  4014
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4015
lemma revcast_rev_ucast [OF refl refl refl]:
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4016
  "cs = [rc, uc] \<Longrightarrow> rc = revcast (word_reverse w) \<Longrightarrow> uc = ucast w \<Longrightarrow>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4017
    rc = word_reverse uc"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4018
  by (metis rev_slice1 revcast_slice1 ucast_slice1 word_size)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4019
45811
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4020
lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))"
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4021
  using revcast_rev_ucast [of "word_reverse w"] by simp
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4022
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4023
lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))"
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4024
  by (fact revcast_rev_ucast [THEN word_rev_gal'])
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4025
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4026
lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)"
f506015ca2dc replace many uses of 'lemmas' with 'lemma';
huffman
parents: 45810
diff changeset
  4027
  by (fact revcast_ucast [THEN word_rev_gal'])
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4028
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4029
65328
2510b0ce28da misc tuning and modernization;
wenzelm
parents: 65268
diff changeset
  4030
text "linking revcast and cast via shift"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4031
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4032
lemmas wsst_TYs = source_size target_size word_size
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4033
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4034
lemmas sym_notr =
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4035
  not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4036
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4037
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4038
subsection \<open>Split and cat\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4039
40827
abbc05c20e24 code preprocessor setup for numerals on word type;
haftmann
parents: 39910
diff changeset
  4040
lemmas word_split_bin' = word_split_def
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4041
lemmas word_cat_bin' = word_cat_eq
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4042
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4043
\<comment> \<open>this odd result is analogous to \<open>ucast_id\<close>,
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4044
      result to the length given by the result type\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4045
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4046
lemma word_cat_id: "word_cat a b = b"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72487
diff changeset
  4047
  by transfer (simp add: take_bit_concat_bit_eq)
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4048
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4049
lemma word_cat_split_alt: "\<lbrakk>size w \<le> size u + size v; word_split w = (u,v)\<rbrakk> \<Longrightarrow> word_cat u v = w"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4050
  unfolding word_split_def
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4051
  by (rule bit_word_eqI) (auto simp: bit_word_cat_iff not_less word_size bit_ucast_iff bit_drop_bit_eq)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4052
45604
29cf40fe8daf eliminated obsolete "standard";
wenzelm
parents: 45550
diff changeset
  4053
lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4054
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4055
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4056
subsubsection \<open>Split and slice\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4057
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4058
lemma split_slices:
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4059
  assumes "word_split w = (u, v)"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4060
  shows "u = slice (size v) w \<and> v = slice 0 w"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4061
  unfolding word_size
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4062
proof (intro conjI)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4063
  have \<section>: "\<And>n. \<lbrakk>ucast (drop_bit LENGTH('b) w) = u; LENGTH('c) < LENGTH('b)\<rbrakk> \<Longrightarrow> \<not> bit u n"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4064
    by (metis bit_take_bit_iff bit_word_of_int_iff diff_is_0_eq' drop_bit_take_bit less_imp_le less_nat_zero_code of_int_uint unsigned_drop_bit_eq)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4065
  show "u = slice LENGTH('b) w"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4066
  proof (rule bit_word_eqI)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4067
    show "bit u n = bit ((slice LENGTH('b) w)::'a word) n" if "n < LENGTH('a)" for n
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4068
      using assms bit_imp_le_length
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4069
      unfolding word_split_def bit_slice_iff
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4070
      by (fastforce simp: \<section> ac_simps word_size bit_ucast_iff bit_drop_bit_eq)
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4071
  qed
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4072
  show "v = slice 0 w"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4073
    by (metis Pair_inject assms ucast_slice word_split_bin')
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4074
qed
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4075
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4076
45816
6a04efd99f25 replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents: 45811
diff changeset
  4077
lemma slice_cat1 [OF refl]:
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4078
  "\<lbrakk>wc = word_cat a b; size a + size b \<le> size wc\<rbrakk> \<Longrightarrow> slice (size b) wc = a"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4079
  by (rule bit_word_eqI) (auto simp: bit_slice_iff bit_word_cat_iff word_size)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4080
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4081
lemmas slice_cat2 = trans [OF slice_id word_cat_id]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4082
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4083
lemma cat_slices:
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4084
  "\<lbrakk>a = slice n c; b = slice 0 c; n = size b; size c \<le> size a + size b\<rbrakk> \<Longrightarrow> word_cat a b = c"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4085
  by (rule bit_word_eqI) (auto simp: bit_slice_iff bit_word_cat_iff word_size)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4086
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4087
lemma word_split_cat_alt:
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4088
  assumes "w = word_cat u v" and size: "size u + size v \<le> size w"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4089
  shows "word_split w = (u,v)"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4090
proof -
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4091
  have "ucast ((drop_bit LENGTH('c) (word_cat u v))::'a word) = u" "ucast ((word_cat u v)::'a word) = v"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4092
    using assms
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4093
    by (auto simp: word_size bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff intro: bit_eqI)
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4094
  then show ?thesis
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4095
    by (simp add: assms(1) word_split_bin')
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4096
qed
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4097
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4098
lemma horner_sum_uint_exp_Cons_eq:
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4099
  \<open>horner_sum uint (2 ^ LENGTH('a)) (w # ws) =
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4100
    concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4101
  for ws :: \<open>'a::len word list\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4102
  by (simp add: bintr_uint concat_bit_eq push_bit_eq_mult)
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4103
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4104
lemma bit_horner_sum_uint_exp_iff:
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4105
  \<open>bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \<longleftrightarrow>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4106
    n div LENGTH('a) < length ws \<and> bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4107
  for ws :: \<open>'a::len word list\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4108
proof (induction ws arbitrary: n)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4109
  case Nil
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4110
  then show ?case
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4111
    by simp
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4112
next
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4113
  case (Cons w ws)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4114
  then show ?case
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4115
    by (cases \<open>n \<ge> LENGTH('a)\<close>)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4116
      (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4117
qed
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4118
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4119
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4120
subsection \<open>Rotation\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4121
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4122
lemma word_rotr_word_rotr_eq: \<open>word_rotr m (word_rotr n w) = word_rotr (m + n) w\<close>
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4123
  by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4124
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4125
lemma word_rot_lem: "\<lbrakk>l + k = d + k mod l; n < l\<rbrakk> \<Longrightarrow> ((d + n) mod l) = n" for l::nat
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4126
  by (metis (no_types, lifting) add.commute add.right_neutral add_diff_cancel_left' mod_if mod_mult_div_eq mod_mult_self2 mod_self)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4127
 
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4128
lemma word_rot_rl [simp]: \<open>word_rotl k (word_rotr k v) = v\<close>
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4129
proof (rule bit_word_eqI)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4130
  show "bit (word_rotl k (word_rotr k v)) n = bit v n" if "n < LENGTH('a)" for n
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4131
    using that
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4132
    by (auto simp: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4133
qed
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4134
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4135
lemma word_rot_lr [simp]: \<open>word_rotr k (word_rotl k v) = v\<close>
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4136
proof (rule bit_word_eqI)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4137
  show "bit (word_rotr k (word_rotl k v)) n = bit v n" if "n < LENGTH('a)" for n
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4138
    using that
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4139
    by (auto simp: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split)
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4140
qed
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4141
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4142
lemma word_rot_gal:
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4143
  \<open>word_rotr n v = w \<longleftrightarrow> word_rotl n w = v\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4144
  by auto
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4145
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4146
lemma word_rot_gal':
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4147
  \<open>w = word_rotr n v \<longleftrightarrow> v = word_rotl n w\<close>
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4148
  by auto
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4149
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4150
lemma word_reverse_word_rotl:
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4151
  \<open>word_reverse (word_rotl n w) = word_rotr n (word_reverse w)\<close> (is \<open>?lhs = ?rhs\<close>)
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4152
proof (rule bit_word_eqI)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4153
  fix m
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4154
  assume \<open>m < LENGTH('a)\<close>
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4155
  then have \<open>int (LENGTH('a) - Suc ((m + n) mod LENGTH('a))) =
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4156
    int ((LENGTH('a) + LENGTH('a) - Suc (m + n mod LENGTH('a))) mod LENGTH('a))\<close>
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  4157
    unfolding of_nat_diff of_nat_mod
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4158
    apply (simp add: Suc_le_eq add_less_le_mono of_nat_mod algebra_simps)
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4159
    apply (simp only: mod_diff_left_eq [symmetric, of \<open>int LENGTH('a) * 2\<close>] mod_mult_self1_is_0 diff_0 minus_mod_int_eq)
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4160
    apply (simp add: mod_simps)
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4161
    done
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4162
  then have \<open>LENGTH('a) - Suc ((m + n) mod LENGTH('a)) =
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4163
            (LENGTH('a) + LENGTH('a) - Suc (m + n mod LENGTH('a))) mod LENGTH('a)\<close>
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4164
    by simp
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4165
  with \<open>m < LENGTH('a)\<close> show \<open>bit ?lhs m \<longleftrightarrow> bit ?rhs m\<close>
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4166
    by (simp add: bit_simps)
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4167
qed
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4168
80401
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4169
lemma word_reverse_word_rotr:
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4170
  \<open>word_reverse (word_rotr n w) = word_rotl n (word_reverse w)\<close>
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4171
  by (rule word_eq_reverseI) (simp add: word_reverse_word_rotl)
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4172
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4173
lemma word_rotl_rev:
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4174
  \<open>word_rotl n w = word_reverse (word_rotr n (word_reverse w))\<close>
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4175
  by (simp add: word_reverse_word_rotr)
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4176
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4177
lemma word_rotr_rev:
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4178
  \<open>word_rotr n w = word_reverse (word_rotl n (word_reverse w))\<close>
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4179
  by (simp add: word_reverse_word_rotl)
31bf95336f16 dropped references to theorems from transitional theory Divides.thy
haftmann
parents: 79950
diff changeset
  4180
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4181
lemma word_roti_0 [simp]: "word_roti 0 w = w"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  4182
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4183
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4184
lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)"
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4185
  by (rule bit_word_eqI)
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4186
    (simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4187
67118
ccab07d1196c more simplification rules
haftmann
parents: 66912
diff changeset
  4188
lemma word_roti_conv_mod':
ccab07d1196c more simplification rules
haftmann
parents: 66912
diff changeset
  4189
  "word_roti n w = word_roti (n mod int (size w)) w"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  4190
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4191
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4192
lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size]
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4193
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4194
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4195
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4196
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4197
subsubsection \<open>"Word rotation commutes with bit-wise operations\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4198
67408
4a4c14b24800 prefer formal comments;
wenzelm
parents: 67399
diff changeset
  4199
\<comment> \<open>using locale to not pollute lemma namespace\<close>
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4200
locale word_rotate
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4201
begin
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4202
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4203
context
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4204
  includes bit_operations_syntax
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4205
begin
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4206
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4207
lemma word_rot_logs:
71149
a7d1fb0c9e16 proper prefix syntax
haftmann
parents: 70901
diff changeset
  4208
  "word_rotl n (NOT v) = NOT (word_rotl n v)"
a7d1fb0c9e16 proper prefix syntax
haftmann
parents: 70901
diff changeset
  4209
  "word_rotr n (NOT v) = NOT (word_rotr n v)"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4210
  "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4211
  "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4212
  "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4213
  "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4214
  "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y"
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4215
  "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4216
  by (rule bit_word_eqI, auto simp: bit_word_rotl_iff bit_word_rotr_iff bit_and_iff bit_or_iff bit_xor_iff bit_not_iff algebra_simps not_le)+
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4217
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4218
end
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4219
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4220
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4221
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4222
lemmas word_rot_logs = word_rotate.word_rot_logs
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4223
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4224
lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \<and> word_rotl i 0 = 0"
72088
a36db1c8238e separation of reversed bit lists from other material
haftmann
parents: 72083
diff changeset
  4225
  by transfer simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4226
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4227
lemma word_roti_0' [simp] : "word_roti n 0 = 0"
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  4228
  by transfer simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4229
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72043
diff changeset
  4230
declare word_roti_eq_word_rotr_word_rotl [simp]
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4231
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4232
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4233
subsection \<open>Maximum machine word\<close>
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4234
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4235
context
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4236
  includes bit_operations_syntax
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4237
begin
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4238
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4239
lemma word_int_cases:
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  4240
  fixes x :: "'a::len word"
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  4241
  obtains n where "x = word_of_int n" and "0 \<le> n" and "n < 2^LENGTH('a)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4242
  by (rule that [of \<open>uint x\<close>]) simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4243
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4244
lemma word_nat_cases [cases type: word]:
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4245
  fixes x :: "'a::len word"
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  4246
  obtains n where "x = of_nat n" and "n < 2^LENGTH('a)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4247
  by (rule that [of \<open>unat x\<close>]) simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4248
73788
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4249
lemma max_word_max [intro!]:
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4250
  \<open>n \<le> - 1\<close> for n :: \<open>'a::len word\<close>
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  4251
  by (fact word_order.extremum)
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4252
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  4253
lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)"
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4254
  by simp
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4255
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  4256
lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0"
71957
3e162c63371a build bit operations on word on library theory on bit operations
haftmann
parents: 71955
diff changeset
  4257
  by (fact word_exp_length_eq_0)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4258
73788
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4259
lemma max_word_wrap: 
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4260
  \<open>x + 1 = 0 \<Longrightarrow> x = - 1\<close> for x :: \<open>'a::len word\<close>
71946
4d4079159be7 replaced mere alias by abbreviation
haftmann
parents: 71945
diff changeset
  4261
  by (simp add: eq_neg_iff_add_eq_0)
4d4079159be7 replaced mere alias by abbreviation
haftmann
parents: 71945
diff changeset
  4262
73788
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4263
lemma word_and_max:
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4264
  \<open>x AND - 1 = x\<close> for x :: \<open>'a::len word\<close>
71946
4d4079159be7 replaced mere alias by abbreviation
haftmann
parents: 71945
diff changeset
  4265
  by (fact word_log_esimps)
4d4079159be7 replaced mere alias by abbreviation
haftmann
parents: 71945
diff changeset
  4266
73788
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4267
lemma word_or_max:
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4268
  \<open>x OR - 1 = - 1\<close> for x :: \<open>'a::len word\<close>
71946
4d4079159be7 replaced mere alias by abbreviation
haftmann
parents: 71945
diff changeset
  4269
  by (fact word_log_esimps)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4270
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4271
lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  4272
  for x y z :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  4273
  by (fact bit.conj_disj_distrib)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4274
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4275
lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  4276
  for x y z :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  4277
  by (fact bit.disj_conj_distrib)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4278
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4279
lemma word_and_not [simp]: "x AND NOT x = 0"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  4280
  for x :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  4281
  by (fact bit.conj_cancel_right)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4282
73788
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4283
lemma word_or_not [simp]:
35217bf33215 max word moved to Word_Lib in AFP
haftmann
parents: 73682
diff changeset
  4284
  \<open>x OR NOT x = - 1\<close> for x :: \<open>'a::len word\<close>
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  4285
  by (fact bit.disj_cancel_right)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4286
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4287
lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y"
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71953
diff changeset
  4288
  for x y :: "'a::len word"
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72489
diff changeset
  4289
  by (fact bit.xor_def)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4290
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4291
lemma uint_lt_0 [simp]: "uint x < 0 = False"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4292
  by (simp add: linorder_not_less)
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4293
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4294
lemma word_less_1 [simp]: "x < 1 \<longleftrightarrow> x = 0"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4295
  for x :: "'a::len word"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4296
  by (simp add: word_less_nat_alt unat_0_iff)
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4297
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4298
lemma uint_plus_if_size:
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4299
  "uint (x + y) =
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4300
    (if uint x + uint y < 2^size x
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4301
     then uint x + uint y
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4302
     else uint x + uint y - 2^size x)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4303
  by (simp add: take_bit_eq_mod word_size uint_word_of_int_eq uint_plus_if')
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4304
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4305
lemma unat_plus_if_size:
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  4306
  "unat (x + y) =
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4307
    (if unat x + unat y < 2^size x
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4308
     then unat x + unat y
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4309
     else unat x + unat y - 2^size x)"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  4310
  for x y :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4311
  by (simp add: size_word.rep_eq unat_arith_simps)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4312
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4313
lemma word_neq_0_conv: "w \<noteq> 0 \<longleftrightarrow> 0 < w"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4314
  for w :: "'a::len word"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4315
  by (fact word_coorder.not_eq_extremum)
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4316
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4317
lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4318
  for c :: "'a::len word"
55818
d8b2f50705d0 more precise imports;
haftmann
parents: 55817
diff changeset
  4319
  by (fact unat_div)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4320
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4321
lemma uint_sub_if_size:
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4322
  "uint (x - y) =
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4323
    (if uint y \<le> uint x
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4324
     then uint x - uint y
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4325
     else uint x - uint y + 2^size x)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4326
  by (simp add: size_word.rep_eq uint_sub_if')
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4327
72130
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  4328
lemma unat_sub:
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  4329
  \<open>unat (a - b) = unat a - unat b\<close>
9e5862223442 dedicated symbols for code generation, to pave way for generic conversions from and to word
haftmann
parents: 72128
diff changeset
  4330
  if \<open>b \<le> a\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4331
  by (meson that unat_sub_if_size word_le_nat_alt)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4332
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  4333
lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46962
diff changeset
  4334
lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4335
70185
ac1706cdde25 clarified notation
haftmann
parents: 70183
diff changeset
  4336
lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4337
  by simp
72292
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4338
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4339
lemma word_of_int_inj:
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4340
  \<open>(word_of_int x :: 'a::len word) = word_of_int y \<longleftrightarrow> x = y\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4341
  if \<open>0 \<le> x \<and> x < 2 ^ LENGTH('a)\<close> \<open>0 \<le> y \<and> y < 2 ^ LENGTH('a)\<close>
4a58c38b85ff factored out typedef material
haftmann
parents: 72281
diff changeset
  4342
  using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self) 
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4343
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4344
lemma word_le_less_eq: "x \<le> y \<longleftrightarrow> x = y \<or> x < y"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4345
  for x y :: "'z::len word"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4346
  by (auto simp: order_class.le_less)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4347
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4348
lemma mod_plus_cong:
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4349
  fixes b b' :: int
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4350
  assumes 1: "b = b'"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4351
    and 2: "x mod b' = x' mod b'"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4352
    and 3: "y mod b' = y' mod b'"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4353
    and 4: "x' + y' = z'"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4354
  shows "(x + y) mod b = z' mod b'"
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4355
proof -
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  4356
  from 1 2[symmetric] 3[symmetric] 
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  4357
  have "(x + y) mod b = (x' mod b' + y' mod b') mod b'"
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 64243
diff changeset
  4358
    by (simp add: mod_add_eq)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4359
  also have "\<dots> = (x' + y') mod b'"
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 64243
diff changeset
  4360
    by (simp add: mod_add_eq)
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4361
  finally show ?thesis
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4362
    by (simp add: 4)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4363
qed
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4364
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4365
lemma mod_minus_cong:
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4366
  fixes b b' :: int
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4367
  assumes "b = b'"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4368
    and "x mod b' = x' mod b'"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4369
    and "y mod b' = y' mod b'"
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4370
    and "x' - y' = z'"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4371
  shows "(x - y) mod b = z' mod b'"
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4372
  using assms [symmetric] by (auto intro: mod_diff_cong)
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4373
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4374
lemma word_induct_less [case_names zero less]:
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4375
  \<open>P m\<close> if zero: \<open>P 0\<close> and less: \<open>\<And>n. n < m \<Longrightarrow> P n \<Longrightarrow> P (1 + n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4376
  for m :: \<open>'a::len word\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4377
proof -
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4378
  define q where \<open>q = unat m\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4379
  with less have \<open>\<And>n. n < word_of_nat q \<Longrightarrow> P n \<Longrightarrow> P (1 + n)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4380
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4381
  then have \<open>P (word_of_nat q :: 'a word)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4382
  proof (induction q)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4383
    case 0
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4384
    show ?case
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4385
      by (simp add: zero)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4386
  next
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4387
    case (Suc q)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4388
    show ?case
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4389
    proof (cases \<open>1 + word_of_nat q = (0 :: 'a word)\<close>)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4390
      case True
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4391
      then show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4392
        by (simp add: zero)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4393
    next
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4394
      case False
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4395
      then have *: \<open>word_of_nat q < (word_of_nat (Suc q) :: 'a word)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4396
        by (simp add: unatSuc word_less_nat_alt)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4397
      then have **: \<open>n < (1 + word_of_nat q :: 'a word) \<longleftrightarrow> n \<le> (word_of_nat q :: 'a word)\<close> for n
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4398
        by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4399
      have \<open>P (word_of_nat q)\<close>
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4400
        by (simp add: "**" Suc.IH Suc.prems)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4401
      with * have \<open>P (1 + word_of_nat q)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4402
        by (rule Suc.prems)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4403
      then show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4404
        by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4405
    qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4406
  qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4407
  with \<open>q = unat m\<close> show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4408
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4409
qed
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4410
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  4411
lemma word_induct: "P 0 \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P m"
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4412
  for P :: "'a::len word \<Rightarrow> bool"
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  4413
  by (rule word_induct_less)
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4414
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4415
lemma word_induct2 [case_names zero suc, induct type]: "P 0 \<Longrightarrow> (\<And>n. 1 + n \<noteq> 0 \<Longrightarrow> P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P n"
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4416
  for P :: "'b::len word \<Rightarrow> bool"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4417
by (induction rule: word_induct_less; force)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4418
55816
e8dd03241e86 cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents: 55415
diff changeset
  4419
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  4420
subsection \<open>Recursion combinator for words\<close>
46010
ebbc2d5cd720 add section headings
huffman
parents: 46009
diff changeset
  4421
54848
a303daddebbf syntactically tuned
haftmann
parents: 54847
diff changeset
  4422
definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a"
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4423
  where "word_rec forZero forSuc n = rec_nat forZero (forSuc \<circ> of_nat) (unat n)"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4424
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4425
lemma word_rec_0 [simp]: "word_rec z s 0 = z"
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4426
  by (simp add: word_rec_def)
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4427
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4428
lemma word_rec_Suc [simp]: "1 + n \<noteq> 0 \<Longrightarrow> word_rec z s (1 + n) = s n (word_rec z s n)"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  4429
  for n :: "'a::len word"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4430
  by (simp add: unatSuc word_rec_def)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4431
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 65336
diff changeset
  4432
lemma word_rec_Pred: "n \<noteq> 0 \<Longrightarrow> word_rec z s n = s (n - 1) (word_rec z s (n - 1))"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4433
  by (metis add.commute diff_add_cancel word_rec_Suc)
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4434
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4435
lemma word_rec_in: "f (word_rec z (\<lambda>_. f) n) = word_rec (f z) (\<lambda>_. f) n"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  4436
  by (induct n) simp_all
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4437
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67122
diff changeset
  4438
lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \<circ> (+) 1) n"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  4439
  by (induct n) simp_all
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4440
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4441
lemma word_rec_twice:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67122
diff changeset
  4442
  "m \<le> n \<Longrightarrow> word_rec z f n = word_rec (word_rec z f (n - m)) (f \<circ> (+) (n - m)) m"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4443
proof (induction n arbitrary: z f)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4444
  case zero
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4445
  then show ?case
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4446
    by (metis diff_0_right word_le_0_iff word_rec_0)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4447
next
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4448
  case (suc n z f)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4449
  show ?case
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4450
  proof (cases "1 + (n - m) = 0")
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4451
    case True
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4452
    then show ?thesis
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4453
      by (simp add: add_diff_eq)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4454
  next
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4455
    case False
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4456
    then have eq: "1 + n - m = 1 + (n - m)"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4457
      by simp
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4458
    with False have "m \<le> n"
81763
2cf8f8e4c1fd Simplified a lot of messy proofs
paulson <lp15@cam.ac.uk>
parents: 81609
diff changeset
  4459
      using inc_le linorder_not_le suc.prems word_le_minus_mono_left by fastforce
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4460
    with False "suc.hyps" show ?thesis
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4461
      using suc.IH [of "f 0 z" "f \<circ> (+) 1"] 
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4462
      by (simp add: word_rec_in2 eq add.assoc o_def)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4463
  qed
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4464
qed
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4465
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4466
lemma word_rec_id: "word_rec z (\<lambda>_. id) n = z"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4467
  by (induct n) auto
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4468
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4469
lemma word_rec_id_eq: "(\<And>m. m < n \<Longrightarrow> f m = id) \<Longrightarrow> word_rec z f n = z"
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4470
  by (induction n) (auto simp: unatSuc unat_arith_simps(2))
37660
56e3520b68b2 one unified Word theory
haftmann
parents: 36899
diff changeset
  4471
65268
75f2aa8ecb12 misc tuning and modernization;
wenzelm
parents: 64593
diff changeset
  4472
lemma word_rec_max:
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4473
  assumes "\<forall>m\<ge>n. m \<noteq> - 1 \<longrightarrow> f m = id"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4474
  shows "word_rec z f (- 1) = word_rec z f n"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4475
proof -
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4476
  have \<section>: "\<And>m. \<lbrakk>m < - 1 - n\<rbrakk> \<Longrightarrow> (f \<circ> (+) n) m = id"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4477
    using assms
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4478
    by (metis (mono_tags, lifting) add.commute add_diff_cancel_left' comp_apply less_le olen_add_eqv plus_minus_no_overflow word_n1_ge)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4479
  have "word_rec z f (- 1) = word_rec (word_rec z f (- 1 - (- 1 - n))) (f \<circ> (+) (- 1 - (- 1 - n))) (- 1 - n)"
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4480
    by (meson word_n1_ge word_rec_twice)
80777
623d46973cbe More tidying of old proofs
paulson <lp15@cam.ac.uk>
parents: 80401
diff changeset
  4481
  also have "\<dots> = word_rec z f n"
72735
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4482
    by (metis (no_types, lifting) \<section> diff_add_cancel minus_diff_eq uminus_add_conv_diff word_rec_id_eq)
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4483
  finally show ?thesis .
bbe5d3ef2052 Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
paulson <lp15@cam.ac.uk>
parents: 72611
diff changeset
  4484
qed
65336
8e5274fc0093 misc tuning and modernization;
wenzelm
parents: 65328
diff changeset
  4485
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4486
end
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73932
diff changeset
  4487
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  4488
82523
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4489
subsection \<open>Some more naive computations rules\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4490
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4491
lemma drop_bit_of_minus_1_eq [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4492
  \<open>drop_bit n (- 1 :: 'a::len word) = mask (LENGTH('a) - n)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4493
  by (rule bit_word_eqI) (auto simp add: bit_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4494
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4495
context
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4496
  includes bit_operations_syntax
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4497
begin
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4498
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4499
lemma word_cat_eq_push_bit_or:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4500
  \<open>word_cat v w = (push_bit LENGTH('b) (ucast v) OR ucast w :: 'c::len word)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4501
  for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4502
  by transfer (simp add: concat_bit_def ac_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4503
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4504
end
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4505
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4506
context semiring_bit_operations
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4507
begin
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4508
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4509
lemma of_nat_take_bit_numeral_eq [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4510
  \<open>of_nat (take_bit m (numeral n)) = take_bit m (numeral n)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4511
  by (simp add: of_nat_take_bit)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4512
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4513
end
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4514
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4515
context ring_bit_operations
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4516
begin
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4517
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4518
lemma signed_take_bit_of_int:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4519
  \<open>signed_take_bit n (of_int k) = of_int (signed_take_bit n k)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4520
  by (rule bit_eqI) (simp add: bit_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4521
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4522
lemma of_int_signed_take_bit:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4523
  \<open>of_int (signed_take_bit n k) = signed_take_bit n (of_int k)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4524
  by (simp add: signed_take_bit_of_int)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4525
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4526
lemma of_int_take_bit_minus_numeral_eq [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4527
  \<open>of_int (take_bit m (numeral n)) = take_bit m (numeral n)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4528
  \<open>of_int (take_bit m (- numeral n)) = take_bit m (- numeral n)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4529
  by (simp_all add: of_int_take_bit)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4530
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4531
end
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4532
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4533
context
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4534
  includes bit_operations_syntax
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4535
begin
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4536
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4537
lemma concat_bit_numeral_of_one_1 [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4538
  \<open>concat_bit (numeral m) 1 l = 1 OR push_bit (numeral m) l\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4539
  by (rule bit_eqI) (auto simp add: bit_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4540
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4541
lemma concat_bit_of_one_2 [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4542
  \<open>concat_bit n k 1 = set_bit n (take_bit n k)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4543
  by (rule bit_eqI) (auto simp add: bit_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4544
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4545
lemma concat_bit_numeral_of_minus_one_1 [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4546
  \<open>concat_bit (numeral m) (- 1) l = push_bit (numeral m) l OR mask (numeral m)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4547
  by (rule bit_eqI) (auto simp add: bit_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4548
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4549
lemma concat_bit_numeral_of_minus_one_2 [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4550
  \<open>concat_bit (numeral m) k (- 1) = take_bit (numeral m) k OR NOT (mask (numeral m))\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4551
  by (rule bit_eqI) (auto simp add: bit_simps)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4552
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4553
lemma concat_bit_numeral [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4554
  \<open>concat_bit (numeral m) (numeral n) (numeral q) = take_bit (numeral m) (numeral n) OR push_bit (numeral m) (numeral q)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4555
  \<open>concat_bit (numeral m) (- numeral n) (numeral q) = take_bit (numeral m) (- numeral n) OR push_bit (numeral m) (numeral q)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4556
  \<open>concat_bit (numeral m) (numeral n) (- numeral q) = take_bit (numeral m) (numeral n) OR push_bit (numeral m) (- numeral q)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4557
  \<open>concat_bit (numeral m) (- numeral n) (- numeral q) = take_bit (numeral m) (- numeral n) OR push_bit (numeral m) (- numeral q)\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4558
  by (fact concat_bit_def)+
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4559
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4560
end
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4561
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4562
lemma word_cat_0_left [simp]:
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4563
  \<open>word_cat 0 w = ucast w\<close>
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4564
  by (simp add: word_cat_eq)
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4565
e4207dfa36b5 explicit check for computations on word type
haftmann
parents: 82376
diff changeset
  4566
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  4567
subsection \<open>Tool support\<close>
72489
a1366ce41368 early and more complete setup of tools
haftmann
parents: 72488
diff changeset
  4568
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69064
diff changeset
  4569
ML_file \<open>Tools/smt_word.ML\<close>
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 35049
diff changeset
  4570
41060
4199fdcfa3c0 moved smt_word.ML into the directory of the Word library
boehmes
parents: 40827
diff changeset
  4571
end