author | haftmann |
Thu, 16 Jul 2020 04:52:25 +0000 | |
changeset 72042 | 587d4681240c |
parent 72027 | 759532ef0885 |
child 72043 | b8bcdb884651 |
permissions | -rw-r--r-- |
29628 | 1 |
(* Title: HOL/Word/Word.thy |
46124 | 2 |
Author: Jeremy Dawson and Gerwin Klein, NICTA |
24333 | 3 |
*) |
4 |
||
61799 | 5 |
section \<open>A type of finite bit strings\<close> |
24350 | 6 |
|
29628 | 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" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
65363
diff
changeset
|
10 |
"HOL-Library.Boolean_Algebra" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
11 |
"HOL-Library.Bit_Operations" |
70190 | 12 |
Bits_Int |
70192 | 13 |
Bit_Comprehension |
71997 | 14 |
Bit_Lists |
53062
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
51717
diff
changeset
|
15 |
Misc_Typedef |
37660 | 16 |
begin |
17 |
||
61799 | 18 |
subsection \<open>Type definition\<close> |
37660 | 19 |
|
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
|
20 |
quotient_type (overloaded) 'a word = int / \<open>\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\<close> |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
21 |
morphisms rep_word word_of_int by (auto intro!: equivpI reflpI sympI transpI) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
22 |
|
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
|
23 |
lift_definition uint :: \<open>'a::len word \<Rightarrow> int\<close> |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
24 |
is \<open>take_bit LENGTH('a)\<close> . |
37660 | 25 |
|
65268 | 26 |
lemma uint_nonnegative: "0 \<le> uint w" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
27 |
by transfer simp |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
28 |
|
70185 | 29 |
lemma uint_bounded: "uint w < 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
|
30 |
for w :: "'a::len word" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
31 |
by transfer (simp add: take_bit_eq_mod) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
32 |
|
70185 | 33 |
lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = 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
|
34 |
for w :: "'a::len word" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
35 |
using uint_nonnegative uint_bounded by (rule mod_pos_pos_trivial) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
36 |
|
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
37 |
lemma word_uint_eqI: "uint a = uint b \<Longrightarrow> a = b" |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
38 |
by transfer simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
39 |
|
65268 | 40 |
lemma word_uint_eq_iff: "a = b \<longleftrightarrow> uint a = uint b" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
41 |
using word_uint_eqI by auto |
70185 | 42 |
|
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
|
43 |
lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
44 |
by transfer (simp add: take_bit_eq_mod) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
45 |
|
65268 | 46 |
lemma word_of_int_uint: "word_of_int (uint w) = w" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
47 |
by transfer simp |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
48 |
|
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
|
49 |
lemma split_word_all: "(\<And>x::'a::len word. PROP P x) \<equiv> (\<And>x. PROP P (word_of_int x))" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
50 |
proof |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
51 |
fix x :: "'a word" |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
52 |
assume "\<And>x. PROP P (word_of_int x)" |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
53 |
then have "PROP P (word_of_int (uint x))" . |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
54 |
then show "PROP P x" by (simp add: word_of_int_uint) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
55 |
qed |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
56 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
57 |
|
61799 | 58 |
subsection \<open>Type conversions and casting\<close> |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
59 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
60 |
definition sint :: "'a::len word \<Rightarrow> int" |
61799 | 61 |
\<comment> \<open>treats the most-significant-bit as a sign bit\<close> |
70175 | 62 |
where sint_uint: "sint w = sbintrunc (LENGTH('a) - 1) (uint w)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
63 |
|
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
|
64 |
definition unat :: "'a::len word \<Rightarrow> nat" |
65268 | 65 |
where "unat w = nat (uint w)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
66 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
67 |
definition uints :: "nat \<Rightarrow> int set" |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
68 |
\<comment> \<open>the sets of integers representing the words\<close> |
65268 | 69 |
where "uints n = range (bintrunc n)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
70 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
71 |
definition sints :: "nat \<Rightarrow> int set" |
65268 | 72 |
where "sints n = range (sbintrunc (n - 1))" |
73 |
||
74 |
lemma uints_num: "uints n = {i. 0 \<le> i \<and> i < 2 ^ n}" |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
75 |
by (simp add: uints_def range_bintrunc) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
76 |
|
65268 | 77 |
lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \<le> i \<and> i < 2 ^ (n - 1)}" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
78 |
by (simp add: sints_def range_sbintrunc) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
79 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
80 |
definition unats :: "nat \<Rightarrow> nat set" |
65268 | 81 |
where "unats n = {i. i < 2 ^ n}" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
82 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
83 |
definition scast :: "'a::len word \<Rightarrow> 'b::len word" |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
84 |
\<comment> \<open>cast a word to a different length\<close> |
65268 | 85 |
where "scast w = word_of_int (sint w)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
86 |
|
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
|
87 |
definition ucast :: "'a::len word \<Rightarrow> 'b::len word" |
65268 | 88 |
where "ucast w = word_of_int (uint w)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
89 |
|
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
|
90 |
instantiation word :: (len) size |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
91 |
begin |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
92 |
|
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
93 |
definition word_size: "size (w :: 'a word) = LENGTH('a)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
94 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
95 |
instance .. |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
96 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
97 |
end |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
98 |
|
65268 | 99 |
lemma word_size_gt_0 [iff]: "0 < size w" |
100 |
for w :: "'a::len word" |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
101 |
by (simp add: word_size) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
102 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
103 |
lemmas lens_gt_0 = word_size_gt_0 len_gt_0 |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
104 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
105 |
lemma lens_not_0 [iff]: |
71942 | 106 |
\<open>size w \<noteq> 0\<close> for w :: \<open>'a::len word\<close> |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
107 |
by auto |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
108 |
|
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
|
109 |
definition source_size :: "('a::len word \<Rightarrow> 'b) \<Rightarrow> nat" |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
110 |
\<comment> \<open>whether a cast (or other) function is to a longer or shorter length\<close> |
65268 | 111 |
where [code del]: "source_size c = (let arb = undefined; x = c arb in size arb)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
112 |
|
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
|
113 |
definition target_size :: "('a \<Rightarrow> 'b::len word) \<Rightarrow> nat" |
65268 | 114 |
where [code del]: "target_size c = size (c undefined)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
115 |
|
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
|
116 |
definition is_up :: "('a::len word \<Rightarrow> 'b::len word) \<Rightarrow> bool" |
65268 | 117 |
where "is_up c \<longleftrightarrow> source_size c \<le> target_size c" |
118 |
||
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
|
119 |
definition is_down :: "('a::len word \<Rightarrow> 'b::len word) \<Rightarrow> bool" |
65268 | 120 |
where "is_down c \<longleftrightarrow> target_size c \<le> source_size c" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
121 |
|
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
|
122 |
definition of_bl :: "bool list \<Rightarrow> 'a::len word" |
65268 | 123 |
where "of_bl bl = word_of_int (bl_to_bin bl)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
124 |
|
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
|
125 |
definition to_bl :: "'a::len word \<Rightarrow> bool list" |
70185 | 126 |
where "to_bl w = bin_to_bl (LENGTH('a)) (uint w)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
127 |
|
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
|
128 |
definition word_int_case :: "(int \<Rightarrow> 'b) \<Rightarrow> 'a::len word \<Rightarrow> 'b" |
65268 | 129 |
where "word_int_case f w = f (uint w)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
130 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
131 |
translations |
65268 | 132 |
"case x of XCONST of_int y \<Rightarrow> b" \<rightleftharpoons> "CONST word_int_case (\<lambda>y. b) x" |
133 |
"case x of (XCONST of_int :: 'a) y \<Rightarrow> b" \<rightharpoonup> "CONST word_int_case (\<lambda>y. b) x" |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
134 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
135 |
|
61799 | 136 |
subsection \<open>Basic code generation setup\<close> |
55817
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
137 |
|
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
|
138 |
definition Word :: "int \<Rightarrow> 'a::len word" |
65268 | 139 |
where [code_post]: "Word = word_of_int" |
140 |
||
141 |
lemma [code abstype]: "Word (uint w) = w" |
|
55817
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
142 |
by (simp add: Word_def word_of_int_uint) |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
143 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
144 |
declare uint_word_of_int [code abstract] |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
145 |
|
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
|
146 |
instantiation word :: (len) equal |
55817
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
147 |
begin |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
148 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
149 |
definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" |
65268 | 150 |
where "equal_word k l \<longleftrightarrow> HOL.equal (uint k) (uint l)" |
151 |
||
152 |
instance |
|
153 |
by standard (simp add: equal equal_word_def word_uint_eq_iff) |
|
55817
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
154 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
155 |
end |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
156 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
157 |
notation fcomp (infixl "\<circ>>" 60) |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
158 |
notation scomp (infixl "\<circ>\<rightarrow>" 60) |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
159 |
|
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
|
160 |
instantiation word :: ("{len, typerep}") random |
55817
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
161 |
begin |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
162 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
163 |
definition |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
164 |
"random_word i = Random.range i \<circ>\<rightarrow> (\<lambda>k. Pair ( |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
165 |
let j = word_of_int (int_of_integer (integer_of_natural k)) :: 'a word |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
166 |
in (j, \<lambda>_::unit. Code_Evaluation.term_of j)))" |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
167 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
168 |
instance .. |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
169 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
170 |
end |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
171 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
172 |
no_notation fcomp (infixl "\<circ>>" 60) |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
173 |
no_notation scomp (infixl "\<circ>\<rightarrow>" 60) |
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
174 |
|
0bc0217387a5
earlier setup of transfer, without dependency on psychodelic interpretations
haftmann
parents:
55816
diff
changeset
|
175 |
|
61799 | 176 |
subsection \<open>Type-definition locale instantiations\<close> |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
177 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
178 |
lemmas uint_0 = uint_nonnegative (* FIXME duplicate *) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
179 |
lemmas uint_lt = uint_bounded (* FIXME duplicate *) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
180 |
lemmas uint_mod_same = uint_idem (* FIXME duplicate *) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
181 |
|
65268 | 182 |
lemma td_ext_uint: |
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
|
183 |
"td_ext (uint :: 'a word \<Rightarrow> int) word_of_int (uints (LENGTH('a::len))) |
70185 | 184 |
(\<lambda>w::int. w mod 2 ^ LENGTH('a))" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
185 |
apply (unfold td_ext_def') |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
186 |
apply transfer |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
187 |
apply (simp add: uints_num take_bit_eq_mod) |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
188 |
done |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
189 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
190 |
interpretation word_uint: |
65268 | 191 |
td_ext |
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
|
192 |
"uint::'a::len word \<Rightarrow> int" |
65268 | 193 |
word_of_int |
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
|
194 |
"uints (LENGTH('a::len))" |
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
|
195 |
"\<lambda>w. w mod 2 ^ LENGTH('a::len)" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
196 |
by (fact td_ext_uint) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
197 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
198 |
lemmas td_uint = word_uint.td_thm |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
199 |
lemmas int_word_uint = word_uint.eq_norm |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
200 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
201 |
lemma td_ext_ubin: |
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
|
202 |
"td_ext (uint :: 'a word \<Rightarrow> int) word_of_int (uints (LENGTH('a::len))) |
70185 | 203 |
(bintrunc (LENGTH('a)))" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
204 |
by (unfold no_bintr_alt1) (fact td_ext_uint) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
205 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
206 |
interpretation word_ubin: |
65268 | 207 |
td_ext |
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
|
208 |
"uint::'a::len word \<Rightarrow> int" |
65268 | 209 |
word_of_int |
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
|
210 |
"uints (LENGTH('a::len))" |
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
|
211 |
"bintrunc (LENGTH('a::len))" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
212 |
by (fact td_ext_ubin) |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
213 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
214 |
|
61799 | 215 |
subsection \<open>Arithmetic operations\<close> |
37660 | 216 |
|
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
|
217 |
lift_definition word_succ :: "'a::len word \<Rightarrow> 'a word" is "\<lambda>x. x + 1" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
218 |
by (auto simp add: bintrunc_mod2p intro: mod_add_cong) |
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
219 |
|
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
|
220 |
lift_definition word_pred :: "'a::len word \<Rightarrow> 'a word" is "\<lambda>x. x - 1" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
221 |
by (auto simp add: bintrunc_mod2p intro: mod_diff_cong) |
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
222 |
|
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
|
223 |
instantiation word :: (len) "{neg_numeral, modulo, comm_monoid_mult, comm_ring}" |
37660 | 224 |
begin |
225 |
||
47387
a0f257197741
remove now-unnecessary type annotations from lift_definition commands
huffman
parents:
47377
diff
changeset
|
226 |
lift_definition zero_word :: "'a word" is "0" . |
a0f257197741
remove now-unnecessary type annotations from lift_definition commands
huffman
parents:
47377
diff
changeset
|
227 |
|
a0f257197741
remove now-unnecessary type annotations from lift_definition commands
huffman
parents:
47377
diff
changeset
|
228 |
lift_definition one_word :: "'a word" is "1" . |
a0f257197741
remove now-unnecessary type annotations from lift_definition commands
huffman
parents:
47377
diff
changeset
|
229 |
|
67399 | 230 |
lift_definition plus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" is "(+)" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
231 |
by (auto simp add: bintrunc_mod2p intro: mod_add_cong) |
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
232 |
|
67399 | 233 |
lift_definition minus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" is "(-)" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
234 |
by (auto simp add: bintrunc_mod2p intro: mod_diff_cong) |
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
235 |
|
47387
a0f257197741
remove now-unnecessary type annotations from lift_definition commands
huffman
parents:
47377
diff
changeset
|
236 |
lift_definition uminus_word :: "'a word \<Rightarrow> 'a word" is uminus |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
237 |
by (auto simp add: bintrunc_mod2p intro: mod_minus_cong) |
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
238 |
|
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68157
diff
changeset
|
239 |
lift_definition times_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" is "(*)" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
240 |
by (auto simp add: bintrunc_mod2p intro: mod_mult_cong) |
37660 | 241 |
|
71950 | 242 |
lift_definition divide_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" |
243 |
is "\<lambda>a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b" |
|
244 |
by simp |
|
245 |
||
246 |
lift_definition modulo_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" |
|
247 |
is "\<lambda>a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b" |
|
248 |
by simp |
|
37660 | 249 |
|
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
250 |
instance |
61169 | 251 |
by standard (transfer, simp add: algebra_simps)+ |
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
252 |
|
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
253 |
end |
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
254 |
|
71950 | 255 |
lemma word_div_def [code]: |
256 |
"a div b = word_of_int (uint a div uint b)" |
|
257 |
by transfer rule |
|
258 |
||
259 |
lemma word_mod_def [code]: |
|
260 |
"a mod b = word_of_int (uint a mod uint b)" |
|
261 |
by transfer rule |
|
262 |
||
70901 | 263 |
quickcheck_generator word |
264 |
constructors: |
|
265 |
"zero_class.zero :: ('a::len) word", |
|
266 |
"numeral :: num \<Rightarrow> ('a::len) word", |
|
267 |
"uminus :: ('a::len) word \<Rightarrow> ('a::len) word" |
|
268 |
||
71950 | 269 |
context |
270 |
includes lifting_syntax |
|
271 |
notes power_transfer [transfer_rule] |
|
272 |
begin |
|
273 |
||
274 |
lemma power_transfer_word [transfer_rule]: |
|
275 |
\<open>(pcr_word ===> (=) ===> pcr_word) (^) (^)\<close> |
|
276 |
by transfer_prover |
|
277 |
||
278 |
end |
|
279 |
||
280 |
||
71951 | 281 |
|
61799 | 282 |
text \<open>Legacy theorems:\<close> |
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
283 |
|
65268 | 284 |
lemma word_arith_wis [code]: |
285 |
shows word_add_def: "a + b = word_of_int (uint a + uint b)" |
|
286 |
and word_sub_wi: "a - b = word_of_int (uint a - uint b)" |
|
287 |
and word_mult_def: "a * b = word_of_int (uint a * uint b)" |
|
288 |
and word_minus_def: "- a = word_of_int (- uint a)" |
|
289 |
and word_succ_alt: "word_succ a = word_of_int (uint a + 1)" |
|
290 |
and word_pred_alt: "word_pred a = word_of_int (uint a - 1)" |
|
291 |
and word_0_wi: "0 = word_of_int 0" |
|
292 |
and word_1_wi: "1 = word_of_int 1" |
|
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
293 |
apply (simp_all flip: plus_word.abs_eq minus_word.abs_eq |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
294 |
times_word.abs_eq uminus_word.abs_eq |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
295 |
zero_word.abs_eq one_word.abs_eq) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
296 |
apply transfer |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
297 |
apply simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
298 |
apply transfer |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
299 |
apply simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
300 |
done |
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
301 |
|
65268 | 302 |
lemma wi_homs: |
303 |
shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" |
|
304 |
and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" |
|
305 |
and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" |
|
306 |
and wi_hom_neg: "- word_of_int a = word_of_int (- a)" |
|
307 |
and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" |
|
308 |
and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" |
|
47374
9475d524bafb
set up and use lift_definition for word operations
huffman
parents:
47372
diff
changeset
|
309 |
by (transfer, simp)+ |
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
310 |
|
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
311 |
lemmas wi_hom_syms = wi_homs [symmetric] |
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
312 |
|
46013 | 313 |
lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi |
46009 | 314 |
|
315 |
lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] |
|
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
316 |
|
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
|
317 |
instance word :: (len) comm_monoid_add .. |
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
|
318 |
|
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
|
319 |
instance word :: (len) semiring_numeral .. |
71950 | 320 |
|
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
321 |
instance word :: (len) comm_ring_1 |
45810 | 322 |
proof |
70185 | 323 |
have *: "0 < LENGTH('a)" by (rule len_gt_0) |
65268 | 324 |
show "(0::'a word) \<noteq> 1" |
325 |
by transfer (use * in \<open>auto simp add: gr0_conv_Suc\<close>) |
|
45810 | 326 |
qed |
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
327 |
|
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
328 |
lemma word_of_nat: "of_nat n = word_of_int (int n)" |
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
329 |
by (induct n) (auto simp add : word_of_int_hom_syms) |
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
330 |
|
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
331 |
lemma word_of_int: "of_int = word_of_int" |
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
332 |
apply (rule ext) |
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
333 |
apply (case_tac x rule: int_diff_cases) |
46013 | 334 |
apply (simp add: word_of_nat wi_hom_sub) |
45545
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
335 |
done |
26aebb8ac9c1
Word.thy: rearrange to instantiate arithmetic classes together with arithmetic operations
huffman
parents:
45544
diff
changeset
|
336 |
|
71950 | 337 |
context |
338 |
includes lifting_syntax |
|
339 |
notes |
|
340 |
transfer_rule_of_bool [transfer_rule] |
|
341 |
transfer_rule_numeral [transfer_rule] |
|
342 |
transfer_rule_of_nat [transfer_rule] |
|
343 |
transfer_rule_of_int [transfer_rule] |
|
344 |
begin |
|
345 |
||
346 |
lemma [transfer_rule]: |
|
347 |
"((=) ===> (pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool)) of_bool of_bool" |
|
348 |
by transfer_prover |
|
349 |
||
350 |
lemma [transfer_rule]: |
|
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
|
351 |
"((=) ===> (pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool)) numeral numeral" |
71950 | 352 |
by transfer_prover |
353 |
||
354 |
lemma [transfer_rule]: |
|
355 |
"((=) ===> pcr_word) int of_nat" |
|
356 |
by transfer_prover |
|
357 |
||
358 |
lemma [transfer_rule]: |
|
359 |
"((=) ===> pcr_word) (\<lambda>k. k) of_int" |
|
360 |
proof - |
|
361 |
have "((=) ===> pcr_word) of_int of_int" |
|
362 |
by transfer_prover |
|
363 |
then show ?thesis by (simp add: id_def) |
|
364 |
qed |
|
365 |
||
366 |
end |
|
367 |
||
368 |
lemma word_of_int_eq: |
|
369 |
"word_of_int = of_int" |
|
370 |
by (rule ext) (transfer, rule) |
|
371 |
||
65268 | 372 |
definition udvd :: "'a::len word \<Rightarrow> 'a::len word \<Rightarrow> bool" (infixl "udvd" 50) |
373 |
where "a udvd b = (\<exists>n\<ge>0. uint b = n * uint a)" |
|
37660 | 374 |
|
71950 | 375 |
context |
376 |
includes lifting_syntax |
|
377 |
begin |
|
378 |
||
379 |
lemma [transfer_rule]: |
|
71958 | 380 |
\<open>(pcr_word ===> (\<longleftrightarrow>)) even ((dvd) 2 :: 'a::len word \<Rightarrow> bool)\<close> |
71950 | 381 |
proof - |
382 |
have even_word_unfold: "even k \<longleftrightarrow> (\<exists>l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \<longleftrightarrow> ?Q") |
|
383 |
for k :: int |
|
384 |
proof |
|
385 |
assume ?P |
|
386 |
then show ?Q |
|
387 |
by auto |
|
388 |
next |
|
389 |
assume ?Q |
|
390 |
then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. |
|
391 |
then have "even (take_bit LENGTH('a) k)" |
|
392 |
by simp |
|
393 |
then show ?P |
|
394 |
by simp |
|
395 |
qed |
|
396 |
show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) |
|
397 |
transfer_prover |
|
398 |
qed |
|
399 |
||
400 |
end |
|
401 |
||
71951 | 402 |
instance word :: (len) semiring_modulo |
403 |
proof |
|
404 |
show "a div b * b + a mod b = a" for a b :: "'a word" |
|
405 |
proof transfer |
|
406 |
fix k l :: int |
|
407 |
define r :: int where "r = 2 ^ LENGTH('a)" |
|
408 |
then have r: "take_bit LENGTH('a) k = k mod r" for k |
|
409 |
by (simp add: take_bit_eq_mod) |
|
410 |
have "k mod r = ((k mod r) div (l mod r) * (l mod r) |
|
411 |
+ (k mod r) mod (l mod r)) mod r" |
|
412 |
by (simp add: div_mult_mod_eq) |
|
413 |
also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r |
|
414 |
+ (k mod r) mod (l mod r)) mod r" |
|
415 |
by (simp add: mod_add_left_eq) |
|
416 |
also have "... = (((k mod r) div (l mod r) * l) mod r |
|
417 |
+ (k mod r) mod (l mod r)) mod r" |
|
418 |
by (simp add: mod_mult_right_eq) |
|
419 |
finally have "k mod r = ((k mod r) div (l mod r) * l |
|
420 |
+ (k mod r) mod (l mod r)) mod r" |
|
421 |
by (simp add: mod_simps) |
|
422 |
with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l |
|
423 |
+ take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" |
|
424 |
by simp |
|
425 |
qed |
|
426 |
qed |
|
427 |
||
428 |
instance word :: (len) semiring_parity |
|
429 |
proof |
|
430 |
show "\<not> 2 dvd (1::'a word)" |
|
431 |
by transfer simp |
|
432 |
show even_iff_mod_2_eq_0: "2 dvd a \<longleftrightarrow> a mod 2 = 0" |
|
433 |
for a :: "'a word" |
|
434 |
by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) |
|
435 |
show "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1" |
|
436 |
for a :: "'a word" |
|
437 |
by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) |
|
438 |
qed |
|
439 |
||
71953 | 440 |
lemma exp_eq_zero_iff: |
441 |
\<open>2 ^ n = (0 :: 'a::len word) \<longleftrightarrow> n \<ge> LENGTH('a)\<close> |
|
442 |
by transfer simp |
|
443 |
||
71958 | 444 |
lemma double_eq_zero_iff: |
445 |
\<open>2 * a = 0 \<longleftrightarrow> a = 0 \<or> a = 2 ^ (LENGTH('a) - Suc 0)\<close> |
|
446 |
for a :: \<open>'a::len word\<close> |
|
447 |
proof - |
|
448 |
define n where \<open>n = LENGTH('a) - Suc 0\<close> |
|
449 |
then have *: \<open>LENGTH('a) = Suc n\<close> |
|
450 |
by simp |
|
451 |
have \<open>a = 0\<close> if \<open>2 * a = 0\<close> and \<open>a \<noteq> 2 ^ (LENGTH('a) - Suc 0)\<close> |
|
452 |
using that by transfer |
|
453 |
(auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) |
|
454 |
moreover have \<open>2 ^ LENGTH('a) = (0 :: 'a word)\<close> |
|
455 |
by transfer simp |
|
456 |
then have \<open>2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\<close> |
|
457 |
by (simp add: *) |
|
458 |
ultimately show ?thesis |
|
459 |
by auto |
|
460 |
qed |
|
461 |
||
45547 | 462 |
|
61799 | 463 |
subsection \<open>Ordering\<close> |
45547 | 464 |
|
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
|
465 |
instantiation word :: (len) linorder |
45547 | 466 |
begin |
467 |
||
71950 | 468 |
lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" |
469 |
is "\<lambda>a b. take_bit LENGTH('a) a \<le> take_bit LENGTH('a) b" |
|
470 |
by simp |
|
471 |
||
472 |
lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" |
|
473 |
is "\<lambda>a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" |
|
474 |
by simp |
|
37660 | 475 |
|
45547 | 476 |
instance |
71950 | 477 |
by (standard; transfer) auto |
45547 | 478 |
|
479 |
end |
|
480 |
||
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
481 |
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
|
482 |
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
|
483 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
484 |
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
|
485 |
by (standard; transfer) simp |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
486 |
|
71950 | 487 |
lemma word_le_def [code]: |
488 |
"a \<le> b \<longleftrightarrow> uint a \<le> uint b" |
|
489 |
by transfer rule |
|
490 |
||
491 |
lemma word_less_def [code]: |
|
492 |
"a < b \<longleftrightarrow> uint a < uint b" |
|
493 |
by transfer rule |
|
494 |
||
71951 | 495 |
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
|
496 |
\<open>a > 0 \<longleftrightarrow> a \<noteq> 0\<close> for a :: \<open>'a::len word\<close> |
71951 | 497 |
by transfer (simp add: less_le) |
498 |
||
499 |
lemma of_nat_word_eq_iff: |
|
500 |
\<open>of_nat m = (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m = take_bit LENGTH('a) n\<close> |
|
501 |
by transfer (simp add: take_bit_of_nat) |
|
502 |
||
503 |
lemma of_nat_word_less_eq_iff: |
|
504 |
\<open>of_nat m \<le> (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close> |
|
505 |
by transfer (simp add: take_bit_of_nat) |
|
506 |
||
507 |
lemma of_nat_word_less_iff: |
|
508 |
\<open>of_nat m < (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close> |
|
509 |
by transfer (simp add: take_bit_of_nat) |
|
510 |
||
511 |
lemma of_nat_word_eq_0_iff: |
|
512 |
\<open>of_nat n = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd n\<close> |
|
513 |
using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) |
|
514 |
||
515 |
lemma of_int_word_eq_iff: |
|
516 |
\<open>of_int k = (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close> |
|
517 |
by transfer rule |
|
518 |
||
519 |
lemma of_int_word_less_eq_iff: |
|
520 |
\<open>of_int k \<le> (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close> |
|
521 |
by transfer rule |
|
522 |
||
523 |
lemma of_int_word_less_iff: |
|
524 |
\<open>of_int k < (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close> |
|
525 |
by transfer rule |
|
526 |
||
527 |
lemma of_int_word_eq_0_iff: |
|
528 |
\<open>of_int k = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd k\<close> |
|
529 |
using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) |
|
530 |
||
65268 | 531 |
definition word_sle :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool" ("(_/ <=s _)" [50, 51] 50) |
532 |
where "a <=s b \<longleftrightarrow> sint a \<le> sint b" |
|
533 |
||
534 |
definition word_sless :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool" ("(_/ <s _)" [50, 51] 50) |
|
535 |
where "x <s y \<longleftrightarrow> x <=s y \<and> x \<noteq> y" |
|
37660 | 536 |
|
537 |
||
61799 | 538 |
subsection \<open>Bit-wise operations\<close> |
37660 | 539 |
|
71951 | 540 |
lemma word_bit_induct [case_names zero even odd]: |
541 |
\<open>P a\<close> if word_zero: \<open>P 0\<close> |
|
542 |
and word_even: \<open>\<And>a. P a \<Longrightarrow> 0 < a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (2 * a)\<close> |
|
543 |
and word_odd: \<open>\<And>a. P a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (1 + 2 * a)\<close> |
|
544 |
for P and a :: \<open>'a::len word\<close> |
|
545 |
proof - |
|
546 |
define m :: nat where \<open>m = LENGTH('a) - 1\<close> |
|
547 |
then have l: \<open>LENGTH('a) = Suc m\<close> |
|
548 |
by simp |
|
549 |
define n :: nat where \<open>n = unat a\<close> |
|
550 |
then have \<open>n < 2 ^ LENGTH('a)\<close> |
|
551 |
by (unfold unat_def) (transfer, simp add: take_bit_eq_mod) |
|
552 |
then have \<open>n < 2 * 2 ^ m\<close> |
|
553 |
by (simp add: l) |
|
554 |
then have \<open>P (of_nat n)\<close> |
|
555 |
proof (induction n rule: nat_bit_induct) |
|
556 |
case zero |
|
557 |
show ?case |
|
558 |
by simp (rule word_zero) |
|
559 |
next |
|
560 |
case (even n) |
|
561 |
then have \<open>n < 2 ^ m\<close> |
|
562 |
by simp |
|
563 |
with even.IH have \<open>P (of_nat n)\<close> |
|
564 |
by simp |
|
565 |
moreover from \<open>n < 2 ^ m\<close> even.hyps have \<open>0 < (of_nat n :: 'a word)\<close> |
|
566 |
by (auto simp add: word_greater_zero_iff of_nat_word_eq_0_iff l) |
|
567 |
moreover from \<open>n < 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close> |
|
568 |
using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>] |
|
569 |
by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l) |
|
570 |
ultimately have \<open>P (2 * of_nat n)\<close> |
|
571 |
by (rule word_even) |
|
572 |
then show ?case |
|
573 |
by simp |
|
574 |
next |
|
575 |
case (odd n) |
|
576 |
then have \<open>Suc n \<le> 2 ^ m\<close> |
|
577 |
by simp |
|
578 |
with odd.IH have \<open>P (of_nat n)\<close> |
|
579 |
by simp |
|
580 |
moreover from \<open>Suc n \<le> 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close> |
|
581 |
using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>] |
|
582 |
by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l) |
|
583 |
ultimately have \<open>P (1 + 2 * of_nat n)\<close> |
|
584 |
by (rule word_odd) |
|
585 |
then show ?case |
|
586 |
by simp |
|
587 |
qed |
|
588 |
moreover have \<open>of_nat (nat (uint a)) = a\<close> |
|
589 |
by transfer simp |
|
590 |
ultimately show ?thesis |
|
591 |
by (simp add: n_def unat_def) |
|
592 |
qed |
|
593 |
||
594 |
lemma bit_word_half_eq: |
|
595 |
\<open>(of_bool b + a * 2) div 2 = a\<close> |
|
596 |
if \<open>a < 2 ^ (LENGTH('a) - Suc 0)\<close> |
|
597 |
for a :: \<open>'a::len word\<close> |
|
598 |
proof (cases \<open>2 \<le> LENGTH('a::len)\<close>) |
|
599 |
case False |
|
600 |
have \<open>of_bool (odd k) < (1 :: int) \<longleftrightarrow> even k\<close> for k :: int |
|
601 |
by auto |
|
602 |
with False that show ?thesis |
|
603 |
by transfer (simp add: eq_iff) |
|
604 |
next |
|
605 |
case True |
|
606 |
obtain n where length: \<open>LENGTH('a) = Suc n\<close> |
|
607 |
by (cases \<open>LENGTH('a)\<close>) simp_all |
|
608 |
show ?thesis proof (cases b) |
|
609 |
case False |
|
610 |
moreover have \<open>a * 2 div 2 = a\<close> |
|
611 |
using that proof transfer |
|
612 |
fix k :: int |
|
613 |
from length have \<open>k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\<close> |
|
614 |
by simp |
|
615 |
moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close> |
|
616 |
with \<open>LENGTH('a) = Suc n\<close> |
|
617 |
have \<open>k mod 2 ^ LENGTH('a) = k mod 2 ^ n\<close> |
|
618 |
by (simp add: take_bit_eq_mod divmod_digit_0) |
|
619 |
ultimately have \<open>take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\<close> |
|
620 |
by (simp add: take_bit_eq_mod) |
|
621 |
with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2) |
|
622 |
= take_bit LENGTH('a) k\<close> |
|
623 |
by simp |
|
624 |
qed |
|
625 |
ultimately show ?thesis |
|
626 |
by simp |
|
627 |
next |
|
628 |
case True |
|
629 |
moreover have \<open>(1 + a * 2) div 2 = a\<close> |
|
630 |
using that proof transfer |
|
631 |
fix k :: int |
|
632 |
from length have \<open>(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\<close> |
|
633 |
using pos_zmod_mult_2 [of \<open>2 ^ n\<close> k] by (simp add: ac_simps) |
|
634 |
moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close> |
|
635 |
with \<open>LENGTH('a) = Suc n\<close> |
|
636 |
have \<open>k mod 2 ^ LENGTH('a) = k mod 2 ^ n\<close> |
|
637 |
by (simp add: take_bit_eq_mod divmod_digit_0) |
|
638 |
ultimately have \<open>take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\<close> |
|
639 |
by (simp add: take_bit_eq_mod) |
|
640 |
with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) |
|
641 |
= take_bit LENGTH('a) k\<close> |
|
642 |
by (auto simp add: take_bit_Suc) |
|
643 |
qed |
|
644 |
ultimately show ?thesis |
|
645 |
by simp |
|
646 |
qed |
|
647 |
qed |
|
648 |
||
649 |
lemma even_mult_exp_div_word_iff: |
|
650 |
\<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> \<not> ( |
|
651 |
m \<le> n \<and> |
|
652 |
n < LENGTH('a) \<and> odd (a div 2 ^ (n - m)))\<close> for a :: \<open>'a::len word\<close> |
|
653 |
by transfer |
|
654 |
(auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff, |
|
655 |
simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int) |
|
656 |
||
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
657 |
instantiation word :: (len) semiring_bits |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
658 |
begin |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
659 |
|
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
660 |
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
|
661 |
is \<open>\<lambda>k n. n < LENGTH('a) \<and> bit k n\<close> |
71951 | 662 |
proof |
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
663 |
fix k l :: int and n :: nat |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
664 |
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
|
665 |
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
|
666 |
proof (cases \<open>n < LENGTH('a)\<close>) |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
667 |
case True |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
668 |
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
|
669 |
by simp |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
670 |
then show ?thesis |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
671 |
by (simp add: bit_take_bit_iff) |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
672 |
next |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
673 |
case False |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
674 |
then show ?thesis |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
675 |
by simp |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
676 |
qed |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
677 |
qed |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
678 |
|
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
679 |
instance proof |
71951 | 680 |
show \<open>P a\<close> if stable: \<open>\<And>a. a div 2 = a \<Longrightarrow> P a\<close> |
681 |
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> |
|
682 |
for P and a :: \<open>'a word\<close> |
|
683 |
proof (induction a rule: word_bit_induct) |
|
684 |
case zero |
|
685 |
have \<open>0 div 2 = (0::'a word)\<close> |
|
686 |
by transfer simp |
|
687 |
with stable [of 0] show ?case |
|
688 |
by simp |
|
689 |
next |
|
690 |
case (even a) |
|
691 |
with rec [of a False] show ?case |
|
692 |
using bit_word_half_eq [of a False] by (simp add: ac_simps) |
|
693 |
next |
|
694 |
case (odd a) |
|
695 |
with rec [of a True] show ?case |
|
696 |
using bit_word_half_eq [of a True] by (simp add: ac_simps) |
|
697 |
qed |
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
698 |
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
|
699 |
by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit) |
71951 | 700 |
show \<open>0 div a = 0\<close> |
701 |
for a :: \<open>'a word\<close> |
|
702 |
by transfer simp |
|
703 |
show \<open>a div 1 = a\<close> |
|
704 |
for a :: \<open>'a word\<close> |
|
705 |
by transfer simp |
|
706 |
show \<open>a mod b div b = 0\<close> |
|
707 |
for a b :: \<open>'a word\<close> |
|
708 |
apply transfer |
|
709 |
apply (simp add: take_bit_eq_mod) |
|
710 |
apply (subst (3) mod_pos_pos_trivial [of _ \<open>2 ^ LENGTH('a)\<close>]) |
|
711 |
apply simp_all |
|
712 |
apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power) |
|
713 |
using pos_mod_bound [of \<open>2 ^ LENGTH('a)\<close>] apply simp |
|
714 |
proof - |
|
715 |
fix aa :: int and ba :: int |
|
716 |
have f1: "\<And>i n. (i::int) mod 2 ^ n = 0 \<or> 0 < i mod 2 ^ n" |
|
717 |
by (metis le_less take_bit_eq_mod take_bit_nonnegative) |
|
718 |
have "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \<and> ba mod 2 ^ len_of (TYPE('a)::'a itself) \<noteq> 0 \<or> aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" |
|
719 |
by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power) |
|
720 |
then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" |
|
721 |
using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound) |
|
722 |
qed |
|
723 |
show \<open>(1 + a) div 2 = a div 2\<close> |
|
724 |
if \<open>even a\<close> |
|
725 |
for a :: \<open>'a word\<close> |
|
71953 | 726 |
using that by transfer |
727 |
(auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE) |
|
71951 | 728 |
show \<open>(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close> |
729 |
for m n :: nat |
|
730 |
by transfer (simp, simp add: exp_div_exp_eq) |
|
731 |
show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" |
|
732 |
for a :: "'a word" and m n :: nat |
|
733 |
apply transfer |
|
734 |
apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) |
|
735 |
apply (simp add: drop_bit_take_bit) |
|
736 |
done |
|
737 |
show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" |
|
738 |
for a :: "'a word" and m n :: nat |
|
739 |
by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) |
|
740 |
show \<open>a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\<close> |
|
741 |
if \<open>m \<le> n\<close> for a :: "'a word" and m n :: nat |
|
742 |
using that apply transfer |
|
743 |
apply (auto simp flip: take_bit_eq_mod) |
|
744 |
apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) |
|
745 |
done |
|
746 |
show \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close> |
|
747 |
for a :: "'a word" and m n :: nat |
|
748 |
by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) |
|
749 |
show \<open>even ((2 ^ m - 1) div (2::'a word) ^ n) \<longleftrightarrow> 2 ^ n = (0::'a word) \<or> m \<le> n\<close> |
|
750 |
for m n :: nat |
|
751 |
by transfer (auto simp add: take_bit_of_mask even_mask_div_iff) |
|
752 |
show \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::'a word) ^ n = 0 \<or> m \<le> n \<and> even (a div 2 ^ (n - m))\<close> |
|
753 |
for a :: \<open>'a word\<close> and m n :: nat |
|
754 |
proof transfer |
|
755 |
show \<open>even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \<longleftrightarrow> |
|
756 |
n < m |
|
757 |
\<or> take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0 |
|
758 |
\<or> (m \<le> n \<and> even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\<close> |
|
759 |
for m n :: nat and k l :: int |
|
760 |
by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult |
|
761 |
simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m]) |
|
762 |
qed |
|
763 |
qed |
|
764 |
||
765 |
end |
|
766 |
||
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
767 |
instantiation word :: (len) semiring_bit_shifts |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
768 |
begin |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
769 |
|
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
770 |
lift_definition push_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
771 |
is push_bit |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
772 |
proof - |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
773 |
show \<open>take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
774 |
if \<open>take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close> for k l :: int and n :: nat |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
775 |
proof - |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
776 |
from that |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
777 |
have \<open>take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
778 |
= take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
779 |
by simp |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
780 |
moreover have \<open>min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
781 |
by simp |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
782 |
ultimately show ?thesis |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
783 |
by (simp add: take_bit_push_bit) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
784 |
qed |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
785 |
qed |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
786 |
|
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
787 |
lift_definition drop_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
788 |
is \<open>\<lambda>n. drop_bit n \<circ> take_bit LENGTH('a)\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
789 |
by (simp add: take_bit_eq_mod) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
790 |
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
791 |
lift_definition take_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close> |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
792 |
is \<open>\<lambda>n. take_bit (min LENGTH('a) n)\<close> |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
793 |
by (simp add: ac_simps) (simp only: flip: take_bit_take_bit) |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
794 |
|
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
795 |
instance proof |
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
796 |
show \<open>push_bit n a = a * 2 ^ n\<close> for n :: nat and a :: \<open>'a word\<close> |
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
797 |
by transfer (simp add: push_bit_eq_mult) |
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
798 |
show \<open>drop_bit n a = a div 2 ^ n\<close> for n :: nat and a :: \<open>'a word\<close> |
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
799 |
by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit) |
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
800 |
show \<open>take_bit n a = a mod 2 ^ n\<close> for n :: nat and a :: \<open>'a word\<close> |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
801 |
by transfer (auto simp flip: take_bit_eq_mod) |
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
802 |
qed |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
803 |
|
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
804 |
end |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
805 |
|
71958 | 806 |
lemma bit_word_eqI: |
807 |
\<open>a = b\<close> if \<open>\<And>n. n \<le> LENGTH('a) \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close> |
|
71990 | 808 |
for a b :: \<open>'a::len word\<close> |
809 |
using that by transfer (auto simp add: nat_less_le bit_eq_iff bit_take_bit_iff) |
|
810 |
||
811 |
lemma bit_imp_le_length: |
|
812 |
\<open>n < LENGTH('a)\<close> if \<open>bit w n\<close> |
|
813 |
for w :: \<open>'a::len word\<close> |
|
814 |
using that by transfer simp |
|
815 |
||
816 |
lemma not_bit_length [simp]: |
|
817 |
\<open>\<not> bit w LENGTH('a)\<close> for w :: \<open>'a::len word\<close> |
|
818 |
by transfer simp |
|
819 |
||
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
820 |
lemma take_bit_length_eq [simp]: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
821 |
\<open>take_bit LENGTH('a) w = w\<close> 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
|
822 |
by transfer simp |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
823 |
|
71990 | 824 |
lemma bit_word_of_int_iff: |
825 |
\<open>bit (word_of_int k :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> bit k n\<close> |
|
826 |
by transfer rule |
|
827 |
||
828 |
lemma bit_uint_iff: |
|
829 |
\<open>bit (uint w) n \<longleftrightarrow> n < LENGTH('a) \<and> bit w n\<close> |
|
830 |
for w :: \<open>'a::len word\<close> |
|
831 |
by transfer (simp add: bit_take_bit_iff) |
|
832 |
||
833 |
lemma bit_sint_iff: |
|
834 |
\<open>bit (sint w) n \<longleftrightarrow> n \<ge> LENGTH('a) \<and> bit w (LENGTH('a) - 1) \<or> bit w n\<close> |
|
835 |
for w :: \<open>'a::len word\<close> |
|
836 |
apply (cases \<open>LENGTH('a)\<close>) |
|
837 |
apply simp |
|
838 |
apply (simp add: sint_uint nth_sbintr not_less bit_uint_iff not_le Suc_le_eq) |
|
839 |
apply (auto simp add: le_less dest: bit_imp_le_length) |
|
840 |
done |
|
841 |
||
842 |
lemma bit_word_ucast_iff: |
|
843 |
\<open>bit (ucast w :: 'b::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> n < LENGTH('b) \<and> bit w n\<close> |
|
844 |
for w :: \<open>'a::len word\<close> |
|
845 |
by (simp add: ucast_def bit_word_of_int_iff bit_uint_iff ac_simps) |
|
846 |
||
847 |
lemma bit_word_scast_iff: |
|
848 |
\<open>bit (scast w :: 'b::len word) n \<longleftrightarrow> |
|
849 |
n < LENGTH('b) \<and> (bit w n \<or> LENGTH('a) \<le> n \<and> bit w (LENGTH('a) - Suc 0))\<close> |
|
850 |
for w :: \<open>'a::len word\<close> |
|
851 |
by (simp add: scast_def bit_word_of_int_iff bit_sint_iff ac_simps) |
|
71958 | 852 |
|
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
|
853 |
definition shiftl1 :: "'a::len word \<Rightarrow> 'a word" |
71986 | 854 |
where "shiftl1 w = word_of_int (2 * uint w)" |
70191 | 855 |
|
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
856 |
lemma shiftl1_eq_mult_2: |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
857 |
\<open>shiftl1 = (*) 2\<close> |
71986 | 858 |
apply (simp add: fun_eq_iff shiftl1_def) |
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
859 |
apply transfer |
71990 | 860 |
apply (simp only: mult_2 take_bit_add) |
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
861 |
apply simp |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
862 |
done |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
863 |
|
71990 | 864 |
lemma bit_shiftl1_iff: |
865 |
\<open>bit (shiftl1 w) n \<longleftrightarrow> 0 < n \<and> n < LENGTH('a) \<and> bit w (n - 1)\<close> |
|
866 |
for w :: \<open>'a::len word\<close> |
|
867 |
by (simp add: shiftl1_eq_mult_2 bit_double_iff exp_eq_zero_iff not_le) (simp add: ac_simps) |
|
868 |
||
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
|
869 |
definition shiftr1 :: "'a::len word \<Rightarrow> 'a word" |
70191 | 870 |
\<comment> \<open>shift right as unsigned or as signed, ie logical or arithmetic\<close> |
871 |
where "shiftr1 w = word_of_int (bin_rest (uint w))" |
|
872 |
||
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
873 |
lemma shiftr1_eq_div_2: |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
874 |
\<open>shiftr1 w = w div 2\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
875 |
apply (simp add: fun_eq_iff shiftr1_def) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
876 |
apply transfer |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
877 |
apply (auto simp add: not_le dest: less_2_cases) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
878 |
done |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
879 |
|
71990 | 880 |
lemma bit_shiftr1_iff: |
881 |
\<open>bit (shiftr1 w) n \<longleftrightarrow> bit w (Suc n)\<close> |
|
882 |
for w :: \<open>'a::len word\<close> |
|
883 |
by (simp add: shiftr1_eq_div_2 bit_Suc) |
|
884 |
||
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
885 |
instantiation word :: (len) ring_bit_operations |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
886 |
begin |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
887 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
888 |
lift_definition not_word :: \<open>'a word \<Rightarrow> 'a word\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
889 |
is not |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
890 |
by (simp add: take_bit_not_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
891 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
892 |
lift_definition and_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
893 |
is \<open>and\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
894 |
by simp |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
895 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
896 |
lift_definition or_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
897 |
is or |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
898 |
by simp |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
899 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
900 |
lift_definition xor_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
901 |
is xor |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
902 |
by simp |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
903 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
904 |
instance proof |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
905 |
fix a b :: \<open>'a word\<close> and n :: nat |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
906 |
show \<open>- a = NOT (a - 1)\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
907 |
by transfer (simp add: minus_eq_not_minus_1) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
908 |
show \<open>bit (NOT a) n \<longleftrightarrow> (2 :: 'a word) ^ n \<noteq> 0 \<and> \<not> bit a n\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
909 |
by transfer (simp add: bit_not_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
910 |
show \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
911 |
by transfer (auto simp add: bit_and_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
912 |
show \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
913 |
by transfer (auto simp add: bit_or_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
914 |
show \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
915 |
by transfer (auto simp add: bit_xor_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
916 |
qed |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
917 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
918 |
end |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
919 |
|
72009 | 920 |
context |
921 |
includes lifting_syntax |
|
922 |
begin |
|
923 |
||
924 |
lemma set_bit_word_transfer: |
|
925 |
\<open>((=) ===> pcr_word ===> pcr_word) set_bit set_bit\<close> |
|
926 |
by (unfold set_bit_def) transfer_prover |
|
927 |
||
928 |
lemma unset_bit_word_transfer: |
|
929 |
\<open>((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\<close> |
|
930 |
by (unfold unset_bit_def) transfer_prover |
|
931 |
||
932 |
lemma flip_bit_word_transfer: |
|
933 |
\<open>((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\<close> |
|
934 |
by (unfold flip_bit_def) transfer_prover |
|
935 |
||
936 |
end |
|
937 |
||
72000 | 938 |
instantiation word :: (len) semiring_bit_syntax |
37660 | 939 |
begin |
940 |
||
65268 | 941 |
definition word_test_bit_def: "test_bit a = bin_nth (uint a)" |
942 |
||
943 |
definition shiftl_def: "w << n = (shiftl1 ^^ n) w" |
|
944 |
||
945 |
definition shiftr_def: "w >> n = (shiftr1 ^^ n) w" |
|
37660 | 946 |
|
71952
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
947 |
lemma test_bit_word_eq: |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
948 |
\<open>test_bit w = bit w\<close> for w :: \<open>'a::len word\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
949 |
apply (simp add: word_test_bit_def fun_eq_iff) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
950 |
apply transfer |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
951 |
apply (simp add: bit_take_bit_iff) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
952 |
done |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
953 |
|
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
954 |
lemma shiftl_word_eq: |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
955 |
\<open>w << n = push_bit n w\<close> for w :: \<open>'a::len word\<close> |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
956 |
by (induction n) (simp_all add: shiftl_def shiftl1_eq_mult_2 push_bit_double) |
2efc5b8c7456
canonical bit shifts for word type, leaving duplicates as they are at the moment
haftmann
parents:
71951
diff
changeset
|
957 |
|
72000 | 958 |
lemma shiftr_word_eq: |
959 |
\<open>w >> n = drop_bit n w\<close> for w :: \<open>'a::len word\<close> |
|
960 |
by (induction n) (simp_all add: shiftr_def shiftr1_eq_div_2 drop_bit_Suc drop_bit_half) |
|
961 |
||
962 |
instance by standard |
|
963 |
(simp_all add: fun_eq_iff test_bit_word_eq shiftl_word_eq shiftr_word_eq) |
|
964 |
||
965 |
end |
|
966 |
||
71990 | 967 |
lemma bit_shiftl_word_iff: |
968 |
\<open>bit (w << m) n \<longleftrightarrow> m \<le> n \<and> n < LENGTH('a) \<and> bit w (n - m)\<close> |
|
969 |
for w :: \<open>'a::len word\<close> |
|
970 |
by (simp add: shiftl_word_eq bit_push_bit_iff exp_eq_zero_iff not_le) |
|
971 |
||
71955 | 972 |
lemma [code]: |
973 |
\<open>push_bit n w = w << n\<close> for w :: \<open>'a::len word\<close> |
|
974 |
by (simp add: shiftl_word_eq) |
|
975 |
||
71990 | 976 |
lemma bit_shiftr_word_iff: |
977 |
\<open>bit (w >> m) n \<longleftrightarrow> bit w (m + n)\<close> |
|
978 |
for w :: \<open>'a::len word\<close> |
|
979 |
by (simp add: shiftr_word_eq bit_drop_bit_eq) |
|
980 |
||
71955 | 981 |
lemma [code]: |
982 |
\<open>drop_bit n w = w >> n\<close> for w :: \<open>'a::len word\<close> |
|
983 |
by (simp add: shiftr_word_eq) |
|
984 |
||
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
985 |
lemma [code]: |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
986 |
\<open>take_bit n a = a AND Bit_Operations.mask n\<close> for a :: \<open>'a::len word\<close> |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
987 |
by (fact take_bit_eq_mask) |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
988 |
|
71955 | 989 |
lemma [code_abbrev]: |
990 |
\<open>push_bit n 1 = (2 :: 'a::len word) ^ n\<close> |
|
991 |
by (fact push_bit_of_1) |
|
992 |
||
65268 | 993 |
lemma [code]: |
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
|
994 |
shows word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" |
65268 | 995 |
and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" |
996 |
and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" |
|
997 |
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
|
998 |
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
|
999 |
|
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
|
1000 |
definition setBit :: "'a::len word \<Rightarrow> nat \<Rightarrow> 'a word" |
72000 | 1001 |
where \<open>setBit w n = Bit_Operations.set_bit n w\<close> |
71990 | 1002 |
|
1003 |
lemma bit_setBit_iff: |
|
1004 |
\<open>bit (setBit w m) n \<longleftrightarrow> (m = n \<and> n < LENGTH('a) \<or> bit w n)\<close> |
|
1005 |
for w :: \<open>'a::len word\<close> |
|
72000 | 1006 |
by (simp add: setBit_def bit_set_bit_iff exp_eq_zero_iff not_le ac_simps) |
71990 | 1007 |
|
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
|
1008 |
definition clearBit :: "'a::len word \<Rightarrow> nat \<Rightarrow> 'a word" |
72000 | 1009 |
where \<open>clearBit w n = unset_bit n w\<close> |
71990 | 1010 |
|
1011 |
lemma bit_clearBit_iff: |
|
1012 |
\<open>bit (clearBit w m) n \<longleftrightarrow> m \<noteq> n \<and> bit w n\<close> |
|
1013 |
for w :: \<open>'a::len word\<close> |
|
72000 | 1014 |
by (simp add: clearBit_def bit_unset_bit_iff ac_simps) |
71990 | 1015 |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1016 |
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
|
1017 |
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
|
1018 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1019 |
lemma even_word_iff [code]: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1020 |
\<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
|
1021 |
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
|
1022 |
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
1023 |
lemma bit_word_iff_drop_bit_and [code]: |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
1024 |
\<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close> for a :: \<open>'a::len word\<close> |
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71958
diff
changeset
|
1025 |
by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1026 |
|
37660 | 1027 |
|
61799 | 1028 |
subsection \<open>Shift operations\<close> |
37660 | 1029 |
|
65268 | 1030 |
definition sshiftr1 :: "'a::len word \<Rightarrow> 'a word" |
1031 |
where "sshiftr1 w = word_of_int (bin_rest (sint w))" |
|
1032 |
||
1033 |
definition bshiftr1 :: "bool \<Rightarrow> 'a::len word \<Rightarrow> 'a word" |
|
1034 |
where "bshiftr1 b w = of_bl (b # butlast (to_bl w))" |
|
1035 |
||
1036 |
definition sshiftr :: "'a::len word \<Rightarrow> nat \<Rightarrow> 'a word" (infixl ">>>" 55) |
|
1037 |
where "w >>> n = (sshiftr1 ^^ n) w" |
|
1038 |
||
1039 |
definition mask :: "nat \<Rightarrow> 'a::len word" |
|
1040 |
where "mask n = (1 << n) - 1" |
|
1041 |
||
37660 | 1042 |
|
61799 | 1043 |
subsection \<open>Rotation\<close> |
37660 | 1044 |
|
65268 | 1045 |
definition rotater1 :: "'a list \<Rightarrow> 'a list" |
1046 |
where "rotater1 ys = |
|
1047 |
(case ys of [] \<Rightarrow> [] | x # xs \<Rightarrow> last ys # butlast ys)" |
|
1048 |
||
1049 |
definition rotater :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
1050 |
where "rotater n = rotater1 ^^ n" |
|
1051 |
||
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
|
1052 |
definition word_rotr :: "nat \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word" |
65268 | 1053 |
where "word_rotr n w = of_bl (rotater n (to_bl w))" |
1054 |
||
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
|
1055 |
definition word_rotl :: "nat \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word" |
65268 | 1056 |
where "word_rotl n w = of_bl (rotate n (to_bl w))" |
1057 |
||
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
|
1058 |
definition word_roti :: "int \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word" |
65268 | 1059 |
where "word_roti i w = |
1060 |
(if i \<ge> 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)" |
|
37660 | 1061 |
|
1062 |
||
61799 | 1063 |
subsection \<open>Split and cat operations\<close> |
37660 | 1064 |
|
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
|
1065 |
definition word_cat :: "'a::len word \<Rightarrow> 'b::len word \<Rightarrow> 'c::len word" |
70185 | 1066 |
where "word_cat a b = word_of_int (bin_cat (uint a) (LENGTH('b)) (uint b))" |
65268 | 1067 |
|
71990 | 1068 |
lemma word_cat_eq: |
1069 |
\<open>(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\<close> |
|
1070 |
for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close> |
|
1071 |
apply (simp add: word_cat_def bin_cat_eq_push_bit_add_take_bit ucast_def) |
|
1072 |
apply transfer apply simp |
|
1073 |
done |
|
1074 |
||
1075 |
lemma bit_word_cat_iff: |
|
1076 |
\<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> |
|
1077 |
for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close> |
|
1078 |
by (auto simp add: word_cat_def bit_word_of_int_iff bin_nth_cat bit_uint_iff not_less bit_imp_le_length) |
|
1079 |
||
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
|
1080 |
definition word_split :: "'a::len word \<Rightarrow> 'b::len word \<times> 'c::len word" |
65268 | 1081 |
where "word_split a = |
70185 | 1082 |
(case bin_split (LENGTH('c)) (uint a) of |
65268 | 1083 |
(u, v) \<Rightarrow> (word_of_int u, word_of_int v))" |
1084 |
||
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
|
1085 |
definition word_rcat :: "'a::len word list \<Rightarrow> 'b::len word" |
70185 | 1086 |
where "word_rcat ws = word_of_int (bin_rcat (LENGTH('a)) (map uint ws))" |
65268 | 1087 |
|
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
|
1088 |
definition word_rsplit :: "'a::len word \<Rightarrow> 'b::len word list" |
70185 | 1089 |
where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" |
65268 | 1090 |
|
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
|
1091 |
abbreviation (input) max_word :: \<open>'a::len word\<close> |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
1092 |
\<comment> \<open>Largest representable machine integer.\<close> |
71946 | 1093 |
where "max_word \<equiv> - 1" |
37660 | 1094 |
|
1095 |
||
61799 | 1096 |
subsection \<open>Theorems about typedefs\<close> |
46010 | 1097 |
|
70185 | 1098 |
lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = sbintrunc (LENGTH('a::len) - 1) bin" |
65268 | 1099 |
by (auto simp: sint_uint word_ubin.eq_norm sbintrunc_bintrunc_lt) |
1100 |
||
70185 | 1101 |
lemma uint_sint: "uint w = bintrunc (LENGTH('a)) (sint w)" |
65328 | 1102 |
for w :: "'a::len word" |
65268 | 1103 |
by (auto simp: sint_uint bintrunc_sbintrunc_le) |
1104 |
||
70185 | 1105 |
lemma bintr_uint: "LENGTH('a) \<le> n \<Longrightarrow> bintrunc 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
|
1106 |
for w :: "'a::len word" |
65268 | 1107 |
apply (subst word_ubin.norm_Rep [symmetric]) |
37660 | 1108 |
apply (simp only: bintrunc_bintrunc_min word_size) |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54854
diff
changeset
|
1109 |
apply (simp add: min.absorb2) |
37660 | 1110 |
done |
1111 |
||
46057 | 1112 |
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
|
1113 |
"LENGTH('a::len) \<le> n \<Longrightarrow> |
46057 | 1114 |
word_of_int (bintrunc n w) = (word_of_int w :: 'a word)" |
65268 | 1115 |
by (auto simp: word_ubin.norm_eq_iff [symmetric] min.absorb1) |
1116 |
||
1117 |
lemma td_ext_sbin: |
|
70185 | 1118 |
"td_ext (sint :: 'a word \<Rightarrow> int) word_of_int (sints (LENGTH('a::len))) |
1119 |
(sbintrunc (LENGTH('a) - 1))" |
|
37660 | 1120 |
apply (unfold td_ext_def' sint_uint) |
1121 |
apply (simp add : word_ubin.eq_norm) |
|
70185 | 1122 |
apply (cases "LENGTH('a)") |
37660 | 1123 |
apply (auto simp add : sints_def) |
1124 |
apply (rule sym [THEN trans]) |
|
65268 | 1125 |
apply (rule word_ubin.Abs_norm) |
37660 | 1126 |
apply (simp only: bintrunc_sbintrunc) |
1127 |
apply (drule sym) |
|
1128 |
apply simp |
|
1129 |
done |
|
1130 |
||
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1131 |
lemma td_ext_sint: |
70185 | 1132 |
"td_ext (sint :: 'a word \<Rightarrow> int) word_of_int (sints (LENGTH('a::len))) |
1133 |
(\<lambda>w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - |
|
1134 |
2 ^ (LENGTH('a) - 1))" |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1135 |
using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) |
37660 | 1136 |
|
67408 | 1137 |
text \<open> |
1138 |
We do \<open>sint\<close> before \<open>sbin\<close>, before \<open>sint\<close> is the user version |
|
1139 |
and interpretations do not produce thm duplicates. I.e. |
|
1140 |
we get the name \<open>word_sint.Rep_eqD\<close>, but not \<open>word_sbin.Req_eqD\<close>, |
|
1141 |
because the latter is the same thm as the former. |
|
1142 |
\<close> |
|
37660 | 1143 |
interpretation word_sint: |
65268 | 1144 |
td_ext |
1145 |
"sint ::'a::len word \<Rightarrow> int" |
|
1146 |
word_of_int |
|
70185 | 1147 |
"sints (LENGTH('a::len))" |
1148 |
"\<lambda>w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - |
|
1149 |
2 ^ (LENGTH('a::len) - 1)" |
|
37660 | 1150 |
by (rule td_ext_sint) |
1151 |
||
1152 |
interpretation word_sbin: |
|
65268 | 1153 |
td_ext |
1154 |
"sint ::'a::len word \<Rightarrow> int" |
|
1155 |
word_of_int |
|
70185 | 1156 |
"sints (LENGTH('a::len))" |
1157 |
"sbintrunc (LENGTH('a::len) - 1)" |
|
37660 | 1158 |
by (rule td_ext_sbin) |
1159 |
||
45604 | 1160 |
lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] |
37660 | 1161 |
|
1162 |
lemmas td_sint = word_sint.td |
|
1163 |
||
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
|
1164 |
lemma to_bl_def': "(to_bl :: 'a::len word \<Rightarrow> bool list) = bin_to_bl (LENGTH('a)) \<circ> uint" |
44762 | 1165 |
by (auto simp: to_bl_def) |
37660 | 1166 |
|
45805 | 1167 |
lemma uints_mod: "uints n = range (\<lambda>w. w mod 2 ^ n)" |
1168 |
by (fact uints_def [unfolded no_bintr_alt1]) |
|
1169 |
||
65268 | 1170 |
lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1171 |
by (induct b, simp_all only: numeral.simps word_of_int_homs) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1172 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1173 |
declare word_numeral_alt [symmetric, code_abbrev] |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1174 |
|
65268 | 1175 |
lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
1176 |
by (simp only: word_numeral_alt wi_hom_neg) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1177 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1178 |
declare word_neg_numeral_alt [symmetric, code_abbrev] |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1179 |
|
45805 | 1180 |
lemma uint_bintrunc [simp]: |
65268 | 1181 |
"uint (numeral bin :: 'a word) = |
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
|
1182 |
bintrunc (LENGTH('a::len)) (numeral bin)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1183 |
unfolding word_numeral_alt by (rule word_ubin.eq_norm) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1184 |
|
65268 | 1185 |
lemma uint_bintrunc_neg [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
|
1186 |
"uint (- numeral bin :: 'a word) = bintrunc (LENGTH('a::len)) (- numeral bin)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1187 |
by (simp only: word_neg_numeral_alt word_ubin.eq_norm) |
37660 | 1188 |
|
45805 | 1189 |
lemma sint_sbintrunc [simp]: |
70185 | 1190 |
"sint (numeral bin :: 'a word) = sbintrunc (LENGTH('a::len) - 1) (numeral bin)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1191 |
by (simp only: word_numeral_alt word_sbin.eq_norm) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1192 |
|
65268 | 1193 |
lemma sint_sbintrunc_neg [simp]: |
70185 | 1194 |
"sint (- numeral bin :: 'a word) = sbintrunc (LENGTH('a::len) - 1) (- numeral bin)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1195 |
by (simp only: word_neg_numeral_alt word_sbin.eq_norm) |
37660 | 1196 |
|
45805 | 1197 |
lemma unat_bintrunc [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
|
1198 |
"unat (numeral bin :: 'a::len word) = nat (bintrunc (LENGTH('a)) (numeral bin))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1199 |
by (simp only: unat_def uint_bintrunc) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1200 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1201 |
lemma unat_bintrunc_neg [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
|
1202 |
"unat (- numeral bin :: 'a::len word) = nat (bintrunc (LENGTH('a)) (- numeral bin))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1203 |
by (simp only: unat_def uint_bintrunc_neg) |
37660 | 1204 |
|
65328 | 1205 |
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
|
1206 |
for v w :: "'a::len word" |
37660 | 1207 |
apply (unfold word_size) |
1208 |
apply (rule word_uint.Rep_eqD) |
|
1209 |
apply (rule box_equals) |
|
1210 |
defer |
|
1211 |
apply (rule word_ubin.norm_Rep)+ |
|
1212 |
apply simp |
|
1213 |
done |
|
1214 |
||
65268 | 1215 |
lemma uint_ge_0 [iff]: "0 \<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
|
1216 |
for x :: "'a::len word" |
45805 | 1217 |
using word_uint.Rep [of x] by (simp add: uints_num) |
1218 |
||
70185 | 1219 |
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
|
1220 |
for x :: "'a::len word" |
45805 | 1221 |
using word_uint.Rep [of x] by (simp add: uints_num) |
1222 |
||
71946 | 1223 |
lemma word_exp_length_eq_0 [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
|
1224 |
\<open>(2 :: 'a::len word) ^ LENGTH('a) = 0\<close> |
71946 | 1225 |
by transfer (simp add: bintrunc_mod2p) |
1226 |
||
70185 | 1227 |
lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \<le> sint x" |
65268 | 1228 |
for x :: "'a::len word" |
45805 | 1229 |
using word_sint.Rep [of x] by (simp add: sints_num) |
1230 |
||
70185 | 1231 |
lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" |
65268 | 1232 |
for x :: "'a::len word" |
45805 | 1233 |
using word_sint.Rep [of x] by (simp add: sints_num) |
37660 | 1234 |
|
65268 | 1235 |
lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1236 |
by (simp add: sign_Pls_ge_0) |
37660 | 1237 |
|
70185 | 1238 |
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
|
1239 |
for x :: "'a::len word" |
45805 | 1240 |
by (simp only: diff_less_0_iff_less uint_lt2p) |
1241 |
||
70185 | 1242 |
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
|
1243 |
for x :: "'a::len word" |
45805 | 1244 |
by (simp only: not_le uint_m2p_neg) |
37660 | 1245 |
|
70185 | 1246 |
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
|
1247 |
for w :: "'a::len word" |
71997 | 1248 |
by (metis bintr_lt2p bintr_uint) |
37660 | 1249 |
|
45805 | 1250 |
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
|
1251 |
by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) |
37660 | 1252 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1253 |
lemma uint_nat: "uint w = int (unat w)" |
65268 | 1254 |
by (auto simp: unat_def) |
1255 |
||
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
|
1256 |
lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" |
65268 | 1257 |
by (simp only: word_numeral_alt int_word_uint) |
1258 |
||
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
|
1259 |
lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" |
65268 | 1260 |
by (simp only: word_neg_numeral_alt int_word_uint) |
1261 |
||
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
|
1262 |
lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" |
37660 | 1263 |
apply (unfold unat_def) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1264 |
apply (clarsimp simp only: uint_numeral) |
37660 | 1265 |
apply (rule nat_mod_distrib [THEN trans]) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1266 |
apply (rule zero_le_numeral) |
37660 | 1267 |
apply (simp_all add: nat_power_eq) |
1268 |
done |
|
1269 |
||
65268 | 1270 |
lemma sint_numeral: |
1271 |
"sint (numeral b :: 'a::len word) = |
|
1272 |
(numeral b + |
|
70185 | 1273 |
2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - |
1274 |
2 ^ (LENGTH('a) - 1)" |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1275 |
unfolding word_numeral_alt by (rule int_word_sint) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1276 |
|
65268 | 1277 |
lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" |
45958 | 1278 |
unfolding word_0_wi .. |
1279 |
||
65268 | 1280 |
lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" |
45958 | 1281 |
unfolding word_1_wi .. |
1282 |
||
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
1283 |
lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
1284 |
by (simp add: wi_hom_syms) |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
1285 |
|
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
|
1286 |
lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" |
65268 | 1287 |
by (simp only: word_numeral_alt) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1288 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1289 |
lemma word_of_int_neg_numeral [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
|
1290 |
"(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" |
65268 | 1291 |
by (simp only: word_numeral_alt wi_hom_syms) |
1292 |
||
1293 |
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
|
1294 |
"word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" |
65268 | 1295 |
by (simp add: word_int_case_def word_uint.eq_norm) |
1296 |
||
1297 |
lemma word_int_split: |
|
1298 |
"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
|
1299 |
(\<forall>i. x = (word_of_int i :: 'b::len word) \<and> 0 \<le> i \<and> i < 2 ^ LENGTH('b) \<longrightarrow> P (f i))" |
71942 | 1300 |
by (auto simp: word_int_case_def word_uint.eq_norm) |
65268 | 1301 |
|
1302 |
lemma word_int_split_asm: |
|
1303 |
"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
|
1304 |
(\<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))" |
71942 | 1305 |
by (auto simp: word_int_case_def word_uint.eq_norm) |
45805 | 1306 |
|
45604 | 1307 |
lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq] |
1308 |
lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq] |
|
37660 | 1309 |
|
65268 | 1310 |
lemma uint_range_size: "0 \<le> uint w \<and> uint w < 2 ^ size w" |
37660 | 1311 |
unfolding word_size by (rule uint_range') |
1312 |
||
65268 | 1313 |
lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \<le> sint w \<and> sint w < 2 ^ (size w - Suc 0)" |
37660 | 1314 |
unfolding word_size by (rule sint_range') |
1315 |
||
65268 | 1316 |
lemma sint_above_size: "2 ^ (size w - 1) \<le> x \<Longrightarrow> sint w < x" |
1317 |
for w :: "'a::len word" |
|
45805 | 1318 |
unfolding word_size by (rule less_le_trans [OF sint_lt]) |
1319 |
||
65268 | 1320 |
lemma sint_below_size: "x \<le> - (2 ^ (size w - 1)) \<Longrightarrow> x \<le> sint w" |
1321 |
for w :: "'a::len word" |
|
45805 | 1322 |
unfolding word_size by (rule order_trans [OF _ sint_ge]) |
37660 | 1323 |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1324 |
|
61799 | 1325 |
subsection \<open>Testing bits\<close> |
46010 | 1326 |
|
65268 | 1327 |
lemma test_bit_eq_iff: "test_bit u = test_bit v \<longleftrightarrow> u = 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
|
1328 |
for u v :: "'a::len word" |
37660 | 1329 |
unfolding word_test_bit_def by (simp add: bin_nth_eq_iff) |
1330 |
||
65268 | 1331 |
lemma test_bit_size [rule_format] : "w !! n \<longrightarrow> n < size 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
|
1332 |
for w :: "'a::len word" |
37660 | 1333 |
apply (unfold word_test_bit_def) |
1334 |
apply (subst word_ubin.norm_Rep [symmetric]) |
|
1335 |
apply (simp only: nth_bintr word_size) |
|
1336 |
apply fast |
|
1337 |
done |
|
1338 |
||
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1339 |
lemma word_eq_iff: "x = y \<longleftrightarrow> (\<forall>n<LENGTH('a). x !! n = y !! n)" (is \<open>?P \<longleftrightarrow> ?Q\<close>) |
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
|
1340 |
for x y :: "'a::len word" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1341 |
proof |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1342 |
assume ?P |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1343 |
then show ?Q |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1344 |
by simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1345 |
next |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1346 |
assume ?Q |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1347 |
then have *: \<open>bit (uint x) n \<longleftrightarrow> bit (uint y) n\<close> if \<open>n < LENGTH('a)\<close> for n |
71949 | 1348 |
using that by (simp add: word_test_bit_def) |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1349 |
show ?P |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1350 |
proof (rule word_uint_eqI, rule bit_eqI, rule iffI) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1351 |
fix n |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1352 |
assume \<open>bit (uint x) n\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1353 |
then have \<open>n < LENGTH('a)\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1354 |
by (simp add: bit_take_bit_iff uint.rep_eq) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1355 |
with * \<open>bit (uint x) n\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1356 |
show \<open>bit (uint y) n\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1357 |
by simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1358 |
next |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1359 |
fix n |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1360 |
assume \<open>bit (uint y) n\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1361 |
then have \<open>n < LENGTH('a)\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1362 |
by (simp add: bit_take_bit_iff uint.rep_eq) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1363 |
with * \<open>bit (uint y) n\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1364 |
show \<open>bit (uint x) n\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1365 |
by simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1366 |
qed |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
1367 |
qed |
46021 | 1368 |
|
65268 | 1369 |
lemma word_eqI: "(\<And>n. n < size u \<longrightarrow> u !! n = v !! n) \<Longrightarrow> u = 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
|
1370 |
for u :: "'a::len word" |
46021 | 1371 |
by (simp add: word_size word_eq_iff) |
37660 | 1372 |
|
65268 | 1373 |
lemma word_eqD: "u = v \<Longrightarrow> u !! x = v !! 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
|
1374 |
for u v :: "'a::len word" |
45805 | 1375 |
by simp |
37660 | 1376 |
|
65268 | 1377 |
lemma test_bit_bin': "w !! n \<longleftrightarrow> n < size w \<and> bin_nth (uint w) n" |
1378 |
by (simp add: word_test_bit_def word_size nth_bintr [symmetric]) |
|
37660 | 1379 |
|
1380 |
lemmas test_bit_bin = test_bit_bin' [unfolded word_size] |
|
1381 |
||
70185 | 1382 |
lemma bin_nth_uint_imp: "bin_nth (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
|
1383 |
for w :: "'a::len word" |
37660 | 1384 |
apply (rule nth_bintr [THEN iffD1, THEN conjunct1]) |
1385 |
apply (subst word_ubin.norm_Rep) |
|
1386 |
apply assumption |
|
1387 |
done |
|
1388 |
||
46057 | 1389 |
lemma bin_nth_sint: |
70185 | 1390 |
"LENGTH('a) \<le> n \<Longrightarrow> |
1391 |
bin_nth (sint w) n = bin_nth (sint w) (LENGTH('a) - 1)" |
|
65268 | 1392 |
for w :: "'a::len word" |
37660 | 1393 |
apply (subst word_sbin.norm_Rep [symmetric]) |
46057 | 1394 |
apply (auto simp add: nth_sbintr) |
37660 | 1395 |
done |
1396 |
||
67408 | 1397 |
\<comment> \<open>type definitions theorem for in terms of equivalent bool list\<close> |
65268 | 1398 |
lemma td_bl: |
1399 |
"type_definition |
|
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
|
1400 |
(to_bl :: 'a::len word \<Rightarrow> bool list) |
65268 | 1401 |
of_bl |
70185 | 1402 |
{bl. length bl = LENGTH('a)}" |
37660 | 1403 |
apply (unfold type_definition_def of_bl_def to_bl_def) |
1404 |
apply (simp add: word_ubin.eq_norm) |
|
1405 |
apply safe |
|
1406 |
apply (drule sym) |
|
1407 |
apply simp |
|
1408 |
done |
|
1409 |
||
1410 |
interpretation word_bl: |
|
65268 | 1411 |
type_definition |
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
|
1412 |
"to_bl :: 'a::len word \<Rightarrow> bool list" |
65268 | 1413 |
of_bl |
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
|
1414 |
"{bl. length bl = LENGTH('a::len)}" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1415 |
by (fact td_bl) |
37660 | 1416 |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
1417 |
lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] |
45538
1fffa81b9b83
eliminated slightly odd Rep' with dynamically-scoped [simplified];
wenzelm
parents:
45529
diff
changeset
|
1418 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1419 |
lemma word_size_bl: "size w = size (to_bl w)" |
65268 | 1420 |
by (auto simp: word_size) |
1421 |
||
1422 |
lemma to_bl_use_of_bl: "to_bl w = bl \<longleftrightarrow> w = of_bl bl \<and> length bl = length (to_bl w)" |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
1423 |
by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) |
37660 | 1424 |
|
65268 | 1425 |
lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" |
1426 |
for x :: "'a::len word" |
|
45805 | 1427 |
unfolding word_bl_Rep' by (rule len_gt_0) |
1428 |
||
65268 | 1429 |
lemma bl_not_Nil [iff]: "to_bl x \<noteq> []" |
1430 |
for x :: "'a::len word" |
|
45805 | 1431 |
by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) |
1432 |
||
65268 | 1433 |
lemma length_bl_neq_0 [iff]: "length (to_bl x) \<noteq> 0" |
1434 |
for x :: "'a::len word" |
|
45805 | 1435 |
by (fact length_bl_gt_0 [THEN gr_implies_not0]) |
37660 | 1436 |
|
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
1437 |
lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" |
37660 | 1438 |
apply (unfold to_bl_def sint_uint) |
1439 |
apply (rule trans [OF _ bl_sbin_sign]) |
|
1440 |
apply simp |
|
1441 |
done |
|
1442 |
||
65268 | 1443 |
lemma of_bl_drop': |
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
|
1444 |
"lend = length bl - LENGTH('a::len) \<Longrightarrow> |
37660 | 1445 |
of_bl (drop lend bl) = (of_bl bl :: 'a word)" |
65268 | 1446 |
by (auto simp: of_bl_def trunc_bl2bin [symmetric]) |
1447 |
||
1448 |
lemma test_bit_of_bl: |
|
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
|
1449 |
"(of_bl bl::'a::len word) !! n = (rev bl ! n \<and> n < LENGTH('a) \<and> n < length bl)" |
65328 | 1450 |
by (auto simp add: of_bl_def word_test_bit_def word_size |
1451 |
word_ubin.eq_norm nth_bintr bin_nth_of_bl) |
|
65268 | 1452 |
|
71990 | 1453 |
lemma bit_of_bl_iff: |
1454 |
\<open>bit (of_bl bs :: 'a word) n \<longleftrightarrow> rev bs ! n \<and> n < LENGTH('a::len) \<and> n < length bs\<close> |
|
1455 |
using test_bit_of_bl [of bs n] by (simp add: test_bit_word_eq) |
|
1456 |
||
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
|
1457 |
lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" |
65268 | 1458 |
by (simp add: of_bl_def) |
37660 | 1459 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1460 |
lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" |
65268 | 1461 |
by (auto simp: word_size to_bl_def) |
37660 | 1462 |
|
1463 |
lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" |
|
65268 | 1464 |
by (simp add: uint_bl word_size) |
1465 |
||
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
|
1466 |
lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" |
65268 | 1467 |
by (auto simp: uint_bl word_ubin.eq_norm word_size) |
37660 | 1468 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1469 |
lemma to_bl_numeral [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
|
1470 |
"to_bl (numeral bin::'a::len word) = |
70185 | 1471 |
bin_to_bl (LENGTH('a)) (numeral bin)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1472 |
unfolding word_numeral_alt by (rule to_bl_of_bin) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1473 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1474 |
lemma to_bl_neg_numeral [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
|
1475 |
"to_bl (- numeral bin::'a::len word) = |
70185 | 1476 |
bin_to_bl (LENGTH('a)) (- numeral bin)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1477 |
unfolding word_neg_numeral_alt by (rule to_bl_of_bin) |
37660 | 1478 |
|
1479 |
lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" |
|
65268 | 1480 |
by (simp add: uint_bl word_size) |
1481 |
||
70185 | 1482 |
lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = 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
|
1483 |
for x :: "'a::len word" |
46011 | 1484 |
by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) |
45604 | 1485 |
|
67408 | 1486 |
\<comment> \<open>naturals\<close> |
37660 | 1487 |
lemma uints_unats: "uints n = int ` unats n" |
1488 |
apply (unfold unats_def uints_num) |
|
1489 |
apply safe |
|
65268 | 1490 |
apply (rule_tac image_eqI) |
1491 |
apply (erule_tac nat_0_le [symmetric]) |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
66808
diff
changeset
|
1492 |
by auto |
37660 | 1493 |
|
1494 |
lemma unats_uints: "unats n = nat ` uints n" |
|
65268 | 1495 |
by (auto simp: uints_unats image_iff) |
1496 |
||
1497 |
lemmas bintr_num = |
|
1498 |
word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b |
|
1499 |
lemmas sbintr_num = |
|
1500 |
word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b |
|
37660 | 1501 |
|
1502 |
lemma num_of_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
|
1503 |
"bintrunc (LENGTH('a::len)) (numeral a) = (numeral b) \<Longrightarrow> |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1504 |
numeral a = (numeral b :: 'a word)" |
46962
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1505 |
unfolding bintr_num by (erule subst, simp) |
37660 | 1506 |
|
1507 |
lemma num_of_sbintr': |
|
70185 | 1508 |
"sbintrunc (LENGTH('a::len) - 1) (numeral a) = (numeral b) \<Longrightarrow> |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1509 |
numeral a = (numeral b :: 'a word)" |
46962
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1510 |
unfolding sbintr_num by (erule subst, simp) |
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1511 |
|
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1512 |
lemma num_abs_bintr: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1513 |
"(numeral x :: 'a word) = |
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
|
1514 |
word_of_int (bintrunc (LENGTH('a::len)) (numeral x))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1515 |
by (simp only: word_ubin.Abs_norm word_numeral_alt) |
46962
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1516 |
|
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1517 |
lemma num_abs_sbintr: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1518 |
"(numeral x :: 'a word) = |
70185 | 1519 |
word_of_int (sbintrunc (LENGTH('a::len) - 1) (numeral x))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1520 |
by (simp only: word_sbin.Abs_norm word_numeral_alt) |
46962
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
1521 |
|
67408 | 1522 |
text \<open> |
1523 |
\<open>cast\<close> -- note, no arg for new length, as it's determined by type of result, |
|
1524 |
thus in \<open>cast w = w\<close>, the type means cast to length of \<open>w\<close>! |
|
1525 |
\<close> |
|
37660 | 1526 |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1527 |
lemma bit_ucast_iff: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1528 |
\<open>Parity.bit (ucast a :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a::len) \<and> Parity.bit a n\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1529 |
by (simp add: ucast_def, transfer) (auto simp add: bit_take_bit_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1530 |
|
37660 | 1531 |
lemma ucast_id: "ucast w = w" |
65268 | 1532 |
by (auto simp: ucast_def) |
37660 | 1533 |
|
1534 |
lemma scast_id: "scast w = w" |
|
65268 | 1535 |
by (auto simp: scast_def) |
37660 | 1536 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1537 |
lemma ucast_bl: "ucast w = of_bl (to_bl w)" |
65268 | 1538 |
by (auto simp: ucast_def of_bl_def uint_bl word_size) |
1539 |
||
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
|
1540 |
lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \<and> n < LENGTH('a))" |
65268 | 1541 |
by (simp add: ucast_def test_bit_bin word_ubin.eq_norm nth_bintr word_size) |
1542 |
(fast elim!: bin_nth_uint_imp) |
|
37660 | 1543 |
|
71958 | 1544 |
context |
1545 |
includes lifting_syntax |
|
1546 |
begin |
|
1547 |
||
1548 |
lemma transfer_rule_mask_word [transfer_rule]: |
|
1549 |
\<open>((=) ===> pcr_word) Bit_Operations.mask Bit_Operations.mask\<close> |
|
1550 |
by (simp only: mask_eq_exp_minus_1 [abs_def]) transfer_prover |
|
1551 |
||
1552 |
end |
|
1553 |
||
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1554 |
lemma ucast_mask_eq: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1555 |
\<open>ucast (Bit_Operations.mask n :: 'b word) = Bit_Operations.mask (min LENGTH('b::len) n)\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1556 |
by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff exp_eq_zero_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1557 |
|
67408 | 1558 |
\<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
|
1559 |
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
|
1560 |
"ucast (numeral w :: 'a::len word) = |
70185 | 1561 |
word_of_int (bintrunc (LENGTH('a)) (numeral w))" |
65268 | 1562 |
by (simp add: ucast_def) |
1563 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1564 |
(* TODO: neg_numeral *) |
37660 | 1565 |
|
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
1566 |
lemma scast_sbintr [simp]: |
65268 | 1567 |
"scast (numeral w ::'a::len word) = |
70185 | 1568 |
word_of_int (sbintrunc (LENGTH('a) - Suc 0) (numeral w))" |
65268 | 1569 |
by (simp add: scast_def) |
37660 | 1570 |
|
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
|
1571 |
lemma source_size: "source_size (c::'a::len word \<Rightarrow> _) = LENGTH('a)" |
46011 | 1572 |
unfolding source_size_def word_size Let_def .. |
1573 |
||
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
|
1574 |
lemma target_size: "target_size (c::_ \<Rightarrow> 'b::len word) = LENGTH('b)" |
46011 | 1575 |
unfolding target_size_def word_size Let_def .. |
1576 |
||
70185 | 1577 |
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
|
1578 |
for c :: "'a::len word \<Rightarrow> 'b::len word" |
65268 | 1579 |
by (simp only: is_down_def source_size target_size) |
1580 |
||
70185 | 1581 |
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
|
1582 |
for c :: "'a::len word \<Rightarrow> 'b::len word" |
65268 | 1583 |
by (simp only: is_up_def source_size target_size) |
37660 | 1584 |
|
45604 | 1585 |
lemmas is_up_down = trans [OF is_up is_down [symmetric]] |
37660 | 1586 |
|
45811 | 1587 |
lemma down_cast_same [OF refl]: "uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc = scast" |
37660 | 1588 |
apply (unfold is_down) |
1589 |
apply safe |
|
1590 |
apply (rule ext) |
|
1591 |
apply (unfold ucast_def scast_def uint_sint) |
|
1592 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
1593 |
apply simp |
|
1594 |
done |
|
1595 |
||
45811 | 1596 |
lemma word_rev_tf: |
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
|
1597 |
"to_bl (of_bl bl::'a::len word) = |
70185 | 1598 |
rev (takefill False (LENGTH('a)) (rev bl))" |
65268 | 1599 |
by (auto simp: of_bl_def uint_bl bl_bin_bl_rtf word_ubin.eq_norm word_size) |
37660 | 1600 |
|
45811 | 1601 |
lemma word_rep_drop: |
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
|
1602 |
"to_bl (of_bl bl::'a::len word) = |
70185 | 1603 |
replicate (LENGTH('a) - length bl) False @ |
1604 |
drop (length bl - LENGTH('a)) bl" |
|
45811 | 1605 |
by (simp add: word_rev_tf takefill_alt rev_take) |
37660 | 1606 |
|
65268 | 1607 |
lemma to_bl_ucast: |
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
|
1608 |
"to_bl (ucast (w::'b::len word) ::'a::len word) = |
70185 | 1609 |
replicate (LENGTH('a) - LENGTH('b)) False @ |
1610 |
drop (LENGTH('b) - LENGTH('a)) (to_bl w)" |
|
37660 | 1611 |
apply (unfold ucast_bl) |
1612 |
apply (rule trans) |
|
1613 |
apply (rule word_rep_drop) |
|
1614 |
apply simp |
|
1615 |
done |
|
1616 |
||
45811 | 1617 |
lemma ucast_up_app [OF refl]: |
65268 | 1618 |
"uc = ucast \<Longrightarrow> source_size uc + n = target_size uc \<Longrightarrow> |
37660 | 1619 |
to_bl (uc w) = replicate n False @ (to_bl w)" |
1620 |
by (auto simp add : source_size target_size to_bl_ucast) |
|
1621 |
||
45811 | 1622 |
lemma ucast_down_drop [OF refl]: |
65268 | 1623 |
"uc = ucast \<Longrightarrow> source_size uc = target_size uc + n \<Longrightarrow> |
37660 | 1624 |
to_bl (uc w) = drop n (to_bl w)" |
1625 |
by (auto simp add : source_size target_size to_bl_ucast) |
|
1626 |
||
45811 | 1627 |
lemma scast_down_drop [OF refl]: |
65268 | 1628 |
"sc = scast \<Longrightarrow> source_size sc = target_size sc + n \<Longrightarrow> |
37660 | 1629 |
to_bl (sc w) = drop n (to_bl w)" |
1630 |
apply (subgoal_tac "sc = ucast") |
|
1631 |
apply safe |
|
1632 |
apply simp |
|
45811 | 1633 |
apply (erule ucast_down_drop) |
1634 |
apply (rule down_cast_same [symmetric]) |
|
37660 | 1635 |
apply (simp add : source_size target_size is_down) |
1636 |
done |
|
1637 |
||
65268 | 1638 |
lemma sint_up_scast [OF refl]: "sc = scast \<Longrightarrow> is_up sc \<Longrightarrow> sint (sc w) = sint w" |
37660 | 1639 |
apply (unfold is_up) |
1640 |
apply safe |
|
1641 |
apply (simp add: scast_def word_sbin.eq_norm) |
|
1642 |
apply (rule box_equals) |
|
1643 |
prefer 3 |
|
1644 |
apply (rule word_sbin.norm_Rep) |
|
1645 |
apply (rule sbintrunc_sbintrunc_l) |
|
1646 |
defer |
|
1647 |
apply (subst word_sbin.norm_Rep) |
|
1648 |
apply (rule refl) |
|
1649 |
apply simp |
|
1650 |
done |
|
1651 |
||
65268 | 1652 |
lemma uint_up_ucast [OF refl]: "uc = ucast \<Longrightarrow> is_up uc \<Longrightarrow> uint (uc w) = uint w" |
37660 | 1653 |
apply (unfold is_up) |
1654 |
apply safe |
|
1655 |
apply (rule bin_eqI) |
|
1656 |
apply (fold word_test_bit_def) |
|
1657 |
apply (auto simp add: nth_ucast) |
|
1658 |
apply (auto simp add: test_bit_bin) |
|
1659 |
done |
|
45811 | 1660 |
|
65268 | 1661 |
lemma ucast_up_ucast [OF refl]: "uc = ucast \<Longrightarrow> is_up uc \<Longrightarrow> ucast (uc w) = ucast w" |
37660 | 1662 |
apply (simp (no_asm) add: ucast_def) |
1663 |
apply (clarsimp simp add: uint_up_ucast) |
|
1664 |
done |
|
65268 | 1665 |
|
1666 |
lemma scast_up_scast [OF refl]: "sc = scast \<Longrightarrow> is_up sc \<Longrightarrow> scast (sc w) = scast w" |
|
37660 | 1667 |
apply (simp (no_asm) add: scast_def) |
1668 |
apply (clarsimp simp add: sint_up_scast) |
|
1669 |
done |
|
65268 | 1670 |
|
1671 |
lemma ucast_of_bl_up [OF refl]: "w = of_bl bl \<Longrightarrow> size bl \<le> size w \<Longrightarrow> ucast w = of_bl bl" |
|
37660 | 1672 |
by (auto simp add : nth_ucast word_size test_bit_of_bl intro!: word_eqI) |
1673 |
||
1674 |
lemmas ucast_up_ucast_id = trans [OF ucast_up_ucast ucast_id] |
|
1675 |
lemmas scast_up_scast_id = trans [OF scast_up_scast scast_id] |
|
1676 |
||
1677 |
lemmas isduu = is_up_down [where c = "ucast", THEN iffD2] |
|
1678 |
lemmas isdus = is_up_down [where c = "scast", THEN iffD2] |
|
1679 |
lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] |
|
1680 |
lemmas scast_down_scast_id = isdus [THEN ucast_up_ucast_id] |
|
1681 |
||
1682 |
lemma up_ucast_surj: |
|
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
|
1683 |
"is_up (ucast :: 'b::len word \<Rightarrow> 'a::len word) \<Longrightarrow> |
65268 | 1684 |
surj (ucast :: 'a word \<Rightarrow> 'b word)" |
1685 |
by (rule surjI) (erule ucast_up_ucast_id) |
|
37660 | 1686 |
|
1687 |
lemma up_scast_surj: |
|
65268 | 1688 |
"is_up (scast :: 'b::len word \<Rightarrow> 'a::len word) \<Longrightarrow> |
1689 |
surj (scast :: 'a word \<Rightarrow> 'b word)" |
|
1690 |
by (rule surjI) (erule scast_up_scast_id) |
|
37660 | 1691 |
|
1692 |
lemma down_scast_inj: |
|
65268 | 1693 |
"is_down (scast :: 'b::len word \<Rightarrow> 'a::len word) \<Longrightarrow> |
1694 |
inj_on (ucast :: 'a word \<Rightarrow> 'b word) A" |
|
37660 | 1695 |
by (rule inj_on_inverseI, erule scast_down_scast_id) |
1696 |
||
1697 |
lemma down_ucast_inj: |
|
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
|
1698 |
"is_down (ucast :: 'b::len word \<Rightarrow> 'a::len word) \<Longrightarrow> |
65268 | 1699 |
inj_on (ucast :: 'a word \<Rightarrow> 'b word) A" |
1700 |
by (rule inj_on_inverseI) (erule ucast_down_ucast_id) |
|
37660 | 1701 |
|
1702 |
lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" |
|
1703 |
by (rule word_bl.Rep_eqD) (simp add: word_rep_drop) |
|
45811 | 1704 |
|
65268 | 1705 |
lemma ucast_down_wi [OF refl]: "uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc (word_of_int x) = word_of_int x" |
46646 | 1706 |
apply (unfold is_down) |
37660 | 1707 |
apply (clarsimp simp add: ucast_def word_ubin.eq_norm) |
1708 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
1709 |
apply (erule bintrunc_bintrunc_ge) |
|
1710 |
done |
|
45811 | 1711 |
|
65268 | 1712 |
lemma ucast_down_no [OF refl]: "uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc (numeral bin) = numeral bin" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1713 |
unfolding word_numeral_alt by clarify (rule ucast_down_wi) |
46646 | 1714 |
|
65268 | 1715 |
lemma ucast_down_bl [OF refl]: "uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc (of_bl bl) = of_bl bl" |
46646 | 1716 |
unfolding of_bl_def by clarify (erule ucast_down_wi) |
37660 | 1717 |
|
1718 |
lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] |
|
1719 |
||
1720 |
lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def |
|
1721 |
||
72000 | 1722 |
lemma bit_last_iff: |
1723 |
\<open>bit w (LENGTH('a) - Suc 0) \<longleftrightarrow> sint w < 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) |
|
1724 |
for w :: \<open>'a::len word\<close> |
|
1725 |
proof - |
|
1726 |
have \<open>?P \<longleftrightarrow> bit (uint w) (LENGTH('a) - Suc 0)\<close> |
|
1727 |
by (simp add: bit_uint_iff) |
|
1728 |
also have \<open>\<dots> \<longleftrightarrow> ?Q\<close> |
|
72010 | 1729 |
by (simp add: sint_uint) |
72000 | 1730 |
finally show ?thesis . |
1731 |
qed |
|
1732 |
||
1733 |
lemma drop_bit_eq_zero_iff_not_bit_last: |
|
1734 |
\<open>drop_bit (LENGTH('a) - Suc 0) w = 0 \<longleftrightarrow> \<not> bit w (LENGTH('a) - Suc 0)\<close> |
|
1735 |
for w :: "'a::len word" |
|
1736 |
apply (cases \<open>LENGTH('a)\<close>) |
|
1737 |
apply simp_all |
|
1738 |
apply (simp add: bit_iff_odd_drop_bit) |
|
1739 |
apply transfer |
|
1740 |
apply (simp add: take_bit_drop_bit) |
|
1741 |
apply (auto simp add: drop_bit_eq_div take_bit_eq_mod min_def) |
|
1742 |
apply (auto elim!: evenE) |
|
1743 |
apply (metis div_exp_eq mod_div_trivial mult.commute nonzero_mult_div_cancel_left power_Suc0_right power_add zero_neq_numeral) |
|
1744 |
done |
|
1745 |
||
37660 | 1746 |
|
61799 | 1747 |
subsection \<open>Word Arithmetic\<close> |
37660 | 1748 |
|
65268 | 1749 |
lemma word_less_alt: "a < b \<longleftrightarrow> uint a < uint b" |
55818 | 1750 |
by (fact word_less_def) |
37660 | 1751 |
|
1752 |
lemma signed_linorder: "class.linorder word_sle word_sless" |
|
65268 | 1753 |
by standard (auto simp: word_sle_def word_sless_def) |
37660 | 1754 |
|
1755 |
interpretation signed: linorder "word_sle" "word_sless" |
|
1756 |
by (rule signed_linorder) |
|
1757 |
||
65268 | 1758 |
lemma udvdI: "0 \<le> n \<Longrightarrow> uint b = n * uint a \<Longrightarrow> a udvd b" |
37660 | 1759 |
by (auto simp: udvd_def) |
1760 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1761 |
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
|
1762 |
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
|
1763 |
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
|
1764 |
lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1765 |
lemmas word_sless_no [simp] = word_sless_def [of "numeral a" "numeral b"] for a b |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1766 |
lemmas word_sle_no [simp] = word_sle_def [of "numeral a" "numeral b"] for a b |
37660 | 1767 |
|
65268 | 1768 |
lemma word_m1_wi: "- 1 = word_of_int (- 1)" |
1769 |
by (simp add: word_neg_numeral_alt [of Num.One]) |
|
37660 | 1770 |
|
46648 | 1771 |
lemma word_0_bl [simp]: "of_bl [] = 0" |
65268 | 1772 |
by (simp add: of_bl_def) |
1773 |
||
1774 |
lemma word_1_bl: "of_bl [True] = 1" |
|
1775 |
by (simp add: of_bl_def bl_to_bin_def) |
|
46648 | 1776 |
|
1777 |
lemma uint_eq_0 [simp]: "uint 0 = 0" |
|
1778 |
unfolding word_0_wi word_ubin.eq_norm by simp |
|
37660 | 1779 |
|
45995
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents:
45958
diff
changeset
|
1780 |
lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" |
46648 | 1781 |
by (simp add: of_bl_def bl_to_bin_rep_False) |
37660 | 1782 |
|
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
|
1783 |
lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" |
65268 | 1784 |
by (simp add: uint_bl word_size bin_to_bl_zero) |
1785 |
||
1786 |
lemma uint_0_iff: "uint x = 0 \<longleftrightarrow> x = 0" |
|
55818 | 1787 |
by (simp add: word_uint_eq_iff) |
1788 |
||
65268 | 1789 |
lemma unat_0_iff: "unat x = 0 \<longleftrightarrow> x = 0" |
1790 |
by (auto simp: unat_def nat_eq_iff uint_0_iff) |
|
1791 |
||
1792 |
lemma unat_0 [simp]: "unat 0 = 0" |
|
1793 |
by (auto simp: unat_def) |
|
1794 |
||
1795 |
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
|
1796 |
for v w :: "'a::len word" |
37660 | 1797 |
apply (unfold word_size) |
1798 |
apply (rule box_equals) |
|
1799 |
defer |
|
1800 |
apply (rule word_uint.Rep_inverse)+ |
|
1801 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
1802 |
apply simp |
|
1803 |
done |
|
1804 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
1805 |
lemmas size_0_same = size_0_same' [unfolded word_size] |
37660 | 1806 |
|
1807 |
lemmas unat_eq_0 = unat_0_iff |
|
1808 |
lemmas unat_eq_zero = unat_0_iff |
|
1809 |
||
65268 | 1810 |
lemma unat_gt_0: "0 < unat x \<longleftrightarrow> x \<noteq> 0" |
1811 |
by (auto simp: unat_0_iff [symmetric]) |
|
37660 | 1812 |
|
45958 | 1813 |
lemma ucast_0 [simp]: "ucast 0 = 0" |
65268 | 1814 |
by (simp add: ucast_def) |
45958 | 1815 |
|
1816 |
lemma sint_0 [simp]: "sint 0 = 0" |
|
65268 | 1817 |
by (simp add: sint_uint) |
45958 | 1818 |
|
1819 |
lemma scast_0 [simp]: "scast 0 = 0" |
|
65268 | 1820 |
by (simp add: scast_def) |
37660 | 1821 |
|
58410
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents:
58061
diff
changeset
|
1822 |
lemma sint_n1 [simp] : "sint (- 1) = - 1" |
65268 | 1823 |
by (simp only: word_m1_wi word_sbin.eq_norm) simp |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
1824 |
|
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
1825 |
lemma scast_n1 [simp]: "scast (- 1) = - 1" |
65268 | 1826 |
by (simp add: scast_def) |
45958 | 1827 |
|
1828 |
lemma uint_1 [simp]: "uint (1::'a::len word) = 1" |
|
71947 | 1829 |
by (simp only: word_1_wi word_ubin.eq_norm) simp |
45958 | 1830 |
|
1831 |
lemma unat_1 [simp]: "unat (1::'a::len word) = 1" |
|
65268 | 1832 |
by (simp add: unat_def) |
45958 | 1833 |
|
1834 |
lemma ucast_1 [simp]: "ucast (1::'a::len word) = 1" |
|
65268 | 1835 |
by (simp add: ucast_def) |
37660 | 1836 |
|
67408 | 1837 |
\<comment> \<open>now, to get the weaker results analogous to \<open>word_div\<close>/\<open>mod_def\<close>\<close> |
37660 | 1838 |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1839 |
|
61799 | 1840 |
subsection \<open>Transferring goals from words to ints\<close> |
37660 | 1841 |
|
65268 | 1842 |
lemma word_ths: |
1843 |
shows word_succ_p1: "word_succ a = a + 1" |
|
1844 |
and word_pred_m1: "word_pred a = a - 1" |
|
1845 |
and word_pred_succ: "word_pred (word_succ a) = a" |
|
1846 |
and word_succ_pred: "word_succ (word_pred a) = a" |
|
1847 |
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
|
1848 |
by (transfer, simp add: algebra_simps)+ |
37660 | 1849 |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
1850 |
lemma uint_cong: "x = y \<Longrightarrow> uint x = uint y" |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
1851 |
by simp |
37660 | 1852 |
|
55818 | 1853 |
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
|
1854 |
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
|
1855 |
shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)" |
70185 | 1856 |
and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" |
1857 |
and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" |
|
1858 |
and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" |
|
1859 |
and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" |
|
1860 |
and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" |
|
1861 |
and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" |
|
1862 |
and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" |
|
55818 | 1863 |
by (simp_all add: word_arith_wis [THEN trans [OF uint_cong int_word_uint]]) |
1864 |
||
1865 |
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
|
1866 |
fixes a b :: "'a::len word" |
70185 | 1867 |
shows "uint (a + b) = bintrunc (LENGTH('a)) (uint a + uint b)" |
1868 |
and "uint (a - b) = bintrunc (LENGTH('a)) (uint a - uint b)" |
|
1869 |
and "uint (a * b) = bintrunc (LENGTH('a)) (uint a * uint b)" |
|
1870 |
and "uint (- a) = bintrunc (LENGTH('a)) (- uint a)" |
|
1871 |
and "uint (word_succ a) = bintrunc (LENGTH('a)) (uint a + 1)" |
|
1872 |
and "uint (word_pred a) = bintrunc (LENGTH('a)) (uint a - 1)" |
|
1873 |
and "uint (0 :: 'a word) = bintrunc (LENGTH('a)) 0" |
|
1874 |
and "uint (1 :: 'a word) = bintrunc (LENGTH('a)) 1" |
|
55818 | 1875 |
by (simp_all add: uint_word_ariths bintrunc_mod2p) |
1876 |
||
1877 |
lemma sint_word_ariths: |
|
1878 |
fixes a b :: "'a::len word" |
|
70185 | 1879 |
shows "sint (a + b) = sbintrunc (LENGTH('a) - 1) (sint a + sint b)" |
1880 |
and "sint (a - b) = sbintrunc (LENGTH('a) - 1) (sint a - sint b)" |
|
1881 |
and "sint (a * b) = sbintrunc (LENGTH('a) - 1) (sint a * sint b)" |
|
1882 |
and "sint (- a) = sbintrunc (LENGTH('a) - 1) (- sint a)" |
|
1883 |
and "sint (word_succ a) = sbintrunc (LENGTH('a) - 1) (sint a + 1)" |
|
1884 |
and "sint (word_pred a) = sbintrunc (LENGTH('a) - 1) (sint a - 1)" |
|
1885 |
and "sint (0 :: 'a word) = sbintrunc (LENGTH('a) - 1) 0" |
|
1886 |
and "sint (1 :: 'a word) = sbintrunc (LENGTH('a) - 1) 1" |
|
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
1887 |
apply (simp_all only: word_sbin.inverse_norm [symmetric]) |
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
1888 |
apply (simp_all add: wi_hom_syms) |
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
1889 |
apply transfer apply simp |
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
1890 |
apply transfer apply simp |
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
1891 |
done |
45604 | 1892 |
|
1893 |
lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] |
|
1894 |
lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] |
|
37660 | 1895 |
|
58410
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents:
58061
diff
changeset
|
1896 |
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
|
1897 |
unfolding word_pred_m1 by simp |
37660 | 1898 |
|
1899 |
lemma succ_pred_no [simp]: |
|
65268 | 1900 |
"word_succ (numeral w) = numeral w + 1" |
1901 |
"word_pred (numeral w) = numeral w - 1" |
|
1902 |
"word_succ (- numeral w) = - numeral w + 1" |
|
1903 |
"word_pred (- numeral w) = - numeral w - 1" |
|
1904 |
by (simp_all add: word_succ_p1 word_pred_m1) |
|
1905 |
||
1906 |
lemma word_sp_01 [simp]: |
|
1907 |
"word_succ (- 1) = 0 \<and> word_succ 0 = 1 \<and> word_pred 0 = - 1 \<and> word_pred 1 = 0" |
|
1908 |
by (simp_all add: word_succ_p1 word_pred_m1) |
|
37660 | 1909 |
|
67408 | 1910 |
\<comment> \<open>alternative approach to lifting arithmetic equalities\<close> |
65268 | 1911 |
lemma word_of_int_Ex: "\<exists>y. x = word_of_int y" |
37660 | 1912 |
by (rule_tac x="uint x" in exI) simp |
1913 |
||
1914 |
||
61799 | 1915 |
subsection \<open>Order on fixed-length words\<close> |
37660 | 1916 |
|
65328 | 1917 |
lemma word_zero_le [simp]: "0 \<le> 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
|
1918 |
for y :: "'a::len word" |
37660 | 1919 |
unfolding word_le_def by auto |
65268 | 1920 |
|
65328 | 1921 |
lemma word_m1_ge [simp] : "word_pred 0 \<ge> y" (* FIXME: delete *) |
71997 | 1922 |
by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 bintr_lt2p) |
65328 | 1923 |
|
1924 |
lemma word_n1_ge [simp]: "y \<le> -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
|
1925 |
for y :: "'a::len word" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
1926 |
by (fact word_order.extremum) |
37660 | 1927 |
|
65268 | 1928 |
lemmas word_not_simps [simp] = |
37660 | 1929 |
word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] |
1930 |
||
65328 | 1931 |
lemma word_gt_0: "0 < y \<longleftrightarrow> 0 \<noteq> 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
|
1932 |
for y :: "'a::len word" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1933 |
by (simp add: less_le) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1934 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
1935 |
lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y |
37660 | 1936 |
|
65328 | 1937 |
lemma word_sless_alt: "a <s b \<longleftrightarrow> sint a < sint b" |
1938 |
by (auto simp add: word_sle_def word_sless_def less_le) |
|
1939 |
||
1940 |
lemma word_le_nat_alt: "a \<le> b \<longleftrightarrow> unat a \<le> unat b" |
|
37660 | 1941 |
unfolding unat_def word_le_def |
1942 |
by (rule nat_le_eq_zle [symmetric]) simp |
|
1943 |
||
65328 | 1944 |
lemma word_less_nat_alt: "a < b \<longleftrightarrow> unat a < unat b" |
37660 | 1945 |
unfolding unat_def word_less_alt |
1946 |
by (rule nat_less_eq_zless [symmetric]) simp |
|
65268 | 1947 |
|
70900 | 1948 |
lemmas unat_mono = word_less_nat_alt [THEN iffD1] |
1949 |
||
1950 |
instance word :: (len) wellorder |
|
1951 |
proof |
|
1952 |
fix P :: "'a word \<Rightarrow> bool" and a |
|
1953 |
assume *: "(\<And>b. (\<And>a. a < b \<Longrightarrow> P a) \<Longrightarrow> P b)" |
|
1954 |
have "wf (measure unat)" .. |
|
1955 |
moreover have "{(a, b :: ('a::len) word). a < b} \<subseteq> measure unat" |
|
1956 |
by (auto simp add: word_less_nat_alt) |
|
1957 |
ultimately have "wf {(a, b :: ('a::len) word). a < b}" |
|
1958 |
by (rule wf_subset) |
|
1959 |
then show "P a" using * |
|
1960 |
by induction blast |
|
1961 |
qed |
|
1962 |
||
65268 | 1963 |
lemma wi_less: |
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
|
1964 |
"(word_of_int n < (word_of_int m :: 'a::len word)) = |
70185 | 1965 |
(n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" |
37660 | 1966 |
unfolding word_less_alt by (simp add: word_uint.eq_norm) |
1967 |
||
65268 | 1968 |
lemma wi_le: |
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
|
1969 |
"(word_of_int n \<le> (word_of_int m :: 'a::len word)) = |
70185 | 1970 |
(n mod 2 ^ LENGTH('a) \<le> m mod 2 ^ LENGTH('a))" |
37660 | 1971 |
unfolding word_le_def by (simp add: word_uint.eq_norm) |
1972 |
||
65328 | 1973 |
lemma udvd_nat_alt: "a udvd b \<longleftrightarrow> (\<exists>n\<ge>0. unat b = n * unat a)" |
37660 | 1974 |
apply (unfold udvd_def) |
1975 |
apply safe |
|
1976 |
apply (simp add: unat_def nat_mult_distrib) |
|
65328 | 1977 |
apply (simp add: uint_nat) |
37660 | 1978 |
apply (rule exI) |
1979 |
apply safe |
|
1980 |
prefer 2 |
|
1981 |
apply (erule notE) |
|
1982 |
apply (rule refl) |
|
1983 |
apply force |
|
1984 |
done |
|
1985 |
||
61941 | 1986 |
lemma udvd_iff_dvd: "x udvd y \<longleftrightarrow> unat x dvd unat y" |
37660 | 1987 |
unfolding dvd_def udvd_nat_alt by force |
1988 |
||
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1989 |
lemma unat_minus_one: |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1990 |
assumes "w \<noteq> 0" |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1991 |
shows "unat (w - 1) = unat w - 1" |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1992 |
proof - |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
1993 |
have "0 \<le> uint w" by (fact uint_nonnegative) |
65328 | 1994 |
moreover from assms have "0 \<noteq> uint w" |
1995 |
by (simp add: uint_0_iff) |
|
1996 |
ultimately have "1 \<le> uint w" |
|
1997 |
by arith |
|
70185 | 1998 |
from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" |
65328 | 1999 |
by arith |
70185 | 2000 |
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
|
2001 |
by (auto intro: mod_pos_pos_trivial) |
70185 | 2002 |
with \<open>1 \<le> uint w\<close> have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2003 |
by auto |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2004 |
then show ?thesis |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
2005 |
by (simp only: unat_def int_word_uint word_arith_wis mod_diff_right_eq) |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2006 |
qed |
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2007 |
|
65328 | 2008 |
lemma measure_unat: "p \<noteq> 0 \<Longrightarrow> unat (p - 1) < unat p" |
37660 | 2009 |
by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) |
65268 | 2010 |
|
45604 | 2011 |
lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] |
2012 |
lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] |
|
37660 | 2013 |
|
70185 | 2014 |
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
|
2015 |
for x :: "'a::len word" and y :: "'b::len word" |
37660 | 2016 |
using uint_ge_0 [of y] uint_lt2p [of x] by arith |
2017 |
||
2018 |
||
61799 | 2019 |
subsection \<open>Conditions for the addition (etc) of two words to overflow\<close> |
37660 | 2020 |
|
65268 | 2021 |
lemma uint_add_lem: |
70185 | 2022 |
"(uint x + uint y < 2 ^ LENGTH('a)) = |
65328 | 2023 |
(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
|
2024 |
for x y :: "'a::len word" |
71997 | 2025 |
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 | 2026 |
|
65268 | 2027 |
lemma uint_mult_lem: |
70185 | 2028 |
"(uint x * uint y < 2 ^ LENGTH('a)) = |
65328 | 2029 |
(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
|
2030 |
for x y :: "'a::len word" |
71997 | 2031 |
by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) |
37660 | 2032 |
|
65328 | 2033 |
lemma uint_sub_lem: "uint x \<ge> uint y \<longleftrightarrow> uint (x - y) = uint x - uint y" |
71997 | 2034 |
by (metis (mono_tags, hide_lams) diff_ge_0_iff_ge mod_pos_pos_trivial of_nat_0_le_iff take_bit_eq_mod uint_nat uint_sub_lt2p word_sub_wi word_ubin.eq_norm) find_theorems uint \<open>- _\<close> |
65328 | 2035 |
|
2036 |
lemma uint_add_le: "uint (x + y) \<le> uint x + uint y" |
|
71997 | 2037 |
unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) |
37660 | 2038 |
|
65328 | 2039 |
lemma uint_sub_ge: "uint (x - y) \<ge> uint x - uint y" |
71997 | 2040 |
unfolding uint_word_ariths by (simp add: int_mod_ge) |
2041 |
||
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2042 |
lemma mod_add_if_z: |
65328 | 2043 |
"x < z \<Longrightarrow> y < z \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> z \<Longrightarrow> |
2044 |
(x + y) mod z = (if x + y < z then x + y else x + y - z)" |
|
2045 |
for x y z :: int |
|
71997 | 2046 |
apply (auto simp add: not_less) |
2047 |
apply (rule antisym) |
|
2048 |
apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) |
|
2049 |
apply (simp only: flip: minus_mod_self2 [of \<open>x + y\<close> z]) |
|
2050 |
apply (rule int_mod_ge) |
|
2051 |
apply simp_all |
|
2052 |
done |
|
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2053 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2054 |
lemma uint_plus_if': |
65328 | 2055 |
"uint (a + b) = |
70185 | 2056 |
(if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b |
2057 |
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
|
2058 |
for a b :: "'a::len word" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2059 |
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
|
2060 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2061 |
lemma mod_sub_if_z: |
65328 | 2062 |
"x < z \<Longrightarrow> y < z \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> z \<Longrightarrow> |
2063 |
(x - y) mod z = (if y \<le> x then x - y else x - y + z)" |
|
2064 |
for x y z :: int |
|
71997 | 2065 |
apply (auto simp add: not_le) |
2066 |
apply (rule antisym) |
|
2067 |
apply (simp only: flip: mod_add_self2 [of \<open>x - y\<close> z]) |
|
2068 |
apply (rule zmod_le_nonneg_dividend) |
|
2069 |
apply simp |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
2070 |
apply (metis add.commute add.right_neutral add_le_cancel_left diff_ge_0_iff_ge int_mod_ge le_less le_less_trans mod_add_self1 not_less) |
71997 | 2071 |
done |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2072 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2073 |
lemma uint_sub_if': |
65328 | 2074 |
"uint (a - b) = |
2075 |
(if uint b \<le> uint a then uint a - uint b |
|
70185 | 2076 |
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
|
2077 |
for a b :: "'a::len word" |
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2078 |
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
|
2079 |
|
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
2080 |
|
61799 | 2081 |
subsection \<open>Definition of \<open>uint_arith\<close>\<close> |
37660 | 2082 |
|
2083 |
lemma word_of_int_inverse: |
|
70185 | 2084 |
"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
|
2085 |
for a :: "'a::len word" |
37660 | 2086 |
apply (erule word_uint.Abs_inverse' [rotated]) |
2087 |
apply (simp add: uints_num) |
|
2088 |
done |
|
2089 |
||
2090 |
lemma uint_split: |
|
70185 | 2091 |
"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
|
2092 |
for x :: "'a::len word" |
37660 | 2093 |
apply (fold word_int_case_def) |
71942 | 2094 |
apply (auto dest!: word_of_int_inverse simp: int_word_uint |
65328 | 2095 |
split: word_int_split) |
37660 | 2096 |
done |
2097 |
||
2098 |
lemma uint_split_asm: |
|
70185 | 2099 |
"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
|
2100 |
for x :: "'a::len word" |
65268 | 2101 |
by (auto dest!: word_of_int_inverse |
71942 | 2102 |
simp: int_word_uint |
65328 | 2103 |
split: uint_split) |
37660 | 2104 |
|
2105 |
lemmas uint_splits = uint_split uint_split_asm |
|
2106 |
||
65268 | 2107 |
lemmas uint_arith_simps = |
37660 | 2108 |
word_le_def word_less_alt |
65268 | 2109 |
word_uint.Rep_inject [symmetric] |
37660 | 2110 |
uint_sub_if' uint_plus_if' |
2111 |
||
70185 | 2112 |
\<comment> \<open>use this to stop, eg. \<open>2 ^ LENGTH(32)\<close> being simplified\<close> |
65268 | 2113 |
lemma power_False_cong: "False \<Longrightarrow> a ^ b = c ^ d" |
37660 | 2114 |
by auto |
2115 |
||
67408 | 2116 |
\<comment> \<open>\<open>uint_arith_tac\<close>: reduce to arithmetic on int, try to solve by arith\<close> |
61799 | 2117 |
ML \<open> |
65268 | 2118 |
fun uint_arith_simpset ctxt = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2119 |
ctxt addsimps @{thms uint_arith_simps} |
37660 | 2120 |
delsimps @{thms word_uint.Rep_inject} |
62390 | 2121 |
|> fold Splitter.add_split @{thms if_split_asm} |
45620
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45604
diff
changeset
|
2122 |
|> fold Simplifier.add_cong @{thms power_False_cong} |
37660 | 2123 |
|
65268 | 2124 |
fun uint_arith_tacs ctxt = |
37660 | 2125 |
let |
2126 |
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
|
2127 |
Arith_Data.arith_tac ctxt n t |
37660 | 2128 |
handle Cooper.COOPER _ => Seq.empty; |
65268 | 2129 |
in |
42793 | 2130 |
[ clarify_tac ctxt 1, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2131 |
full_simp_tac (uint_arith_simpset ctxt) 1, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2132 |
ALLGOALS (full_simp_tac |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2133 |
(put_simpset HOL_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2134 |
|> fold Splitter.add_split @{thms uint_splits} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2135 |
|> fold Simplifier.add_cong @{thms power_False_cong})), |
65268 | 2136 |
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
|
2137 |
ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN |
60754 | 2138 |
REPEAT (eresolve_tac ctxt [conjE] n) THEN |
65268 | 2139 |
REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n |
2140 |
THEN assume_tac ctxt n |
|
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58874
diff
changeset
|
2141 |
THEN assume_tac ctxt n)), |
37660 | 2142 |
TRYALL arith_tac' ] |
2143 |
end |
|
2144 |
||
2145 |
fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) |
|
61799 | 2146 |
\<close> |
37660 | 2147 |
|
65268 | 2148 |
method_setup uint_arith = |
61799 | 2149 |
\<open>Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\<close> |
37660 | 2150 |
"solving word arithmetic via integers and arith" |
2151 |
||
2152 |
||
61799 | 2153 |
subsection \<open>More on overflows and monotonicity\<close> |
37660 | 2154 |
|
65328 | 2155 |
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
|
2156 |
for x y :: "'a::len word" |
37660 | 2157 |
unfolding word_size by uint_arith |
2158 |
||
2159 |
lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] |
|
2160 |
||
65328 | 2161 |
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
|
2162 |
for x y :: "'a::len word" |
37660 | 2163 |
by uint_arith |
2164 |
||
70185 | 2165 |
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
|
2166 |
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
|
2167 |
by (simp add: ac_simps no_olen_add) |
37660 | 2168 |
|
45604 | 2169 |
lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] |
2170 |
||
2171 |
lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] |
|
2172 |
lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] |
|
2173 |
lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] |
|
37660 | 2174 |
lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] |
2175 |
lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] |
|
45604 | 2176 |
lemmas word_sub_le = word_sub_le_iff [THEN iffD2] |
37660 | 2177 |
|
65328 | 2178 |
lemma word_less_sub1: "x \<noteq> 0 \<Longrightarrow> 1 < x \<longleftrightarrow> 0 < x - 1" |
2179 |
for x :: "'a::len word" |
|
37660 | 2180 |
by uint_arith |
2181 |
||
65328 | 2182 |
lemma word_le_sub1: "x \<noteq> 0 \<Longrightarrow> 1 \<le> x \<longleftrightarrow> 0 \<le> x - 1" |
2183 |
for x :: "'a::len word" |
|
37660 | 2184 |
by uint_arith |
2185 |
||
65328 | 2186 |
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
|
2187 |
for x z :: "'a::len word" |
37660 | 2188 |
by uint_arith |
2189 |
||
65328 | 2190 |
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
|
2191 |
for x z :: "'a::len word" |
37660 | 2192 |
by uint_arith |
2193 |
||
65328 | 2194 |
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
|
2195 |
for x ab c :: "'a::len word" |
37660 | 2196 |
by uint_arith |
2197 |
||
65328 | 2198 |
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
|
2199 |
for x ab c :: "'a::len word" |
37660 | 2200 |
by uint_arith |
2201 |
||
65328 | 2202 |
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
|
2203 |
for a b c :: "'a::len word" |
37660 | 2204 |
by uint_arith |
2205 |
||
65328 | 2206 |
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
|
2207 |
for a b c :: "'a::len word" |
37660 | 2208 |
by uint_arith |
2209 |
||
2210 |
lemmas le_plus = le_plus' [rotated] |
|
2211 |
||
46011 | 2212 |
lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) |
37660 | 2213 |
|
65328 | 2214 |
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
|
2215 |
for x y z :: "'a::len word" |
37660 | 2216 |
by uint_arith |
2217 |
||
65328 | 2218 |
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
|
2219 |
for x y z :: "'a::len word" |
37660 | 2220 |
by uint_arith |
2221 |
||
65328 | 2222 |
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
|
2223 |
for x y z :: "'a::len word" |
37660 | 2224 |
by uint_arith |
2225 |
||
65328 | 2226 |
lemma word_less_minus_mono: "a < c \<Longrightarrow> d < b \<Longrightarrow> a - b < a \<Longrightarrow> c - d < c \<Longrightarrow> a - b < c - d" |
2227 |
for a b c d :: "'a::len word" |
|
37660 | 2228 |
by uint_arith |
2229 |
||
65328 | 2230 |
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
|
2231 |
for x y z :: "'a::len word" |
37660 | 2232 |
by uint_arith |
2233 |
||
65328 | 2234 |
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
|
2235 |
for x y z :: "'a::len word" |
37660 | 2236 |
by uint_arith |
2237 |
||
65268 | 2238 |
lemma word_le_minus_mono: |
65328 | 2239 |
"a \<le> c \<Longrightarrow> d \<le> b \<Longrightarrow> a - b \<le> a \<Longrightarrow> c - d \<le> c \<Longrightarrow> a - b \<le> c - d" |
2240 |
for a b c d :: "'a::len word" |
|
37660 | 2241 |
by uint_arith |
2242 |
||
65328 | 2243 |
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
|
2244 |
for x y y' :: "'a::len word" |
37660 | 2245 |
by uint_arith |
2246 |
||
65328 | 2247 |
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
|
2248 |
for x y y' :: "'a::len word" |
37660 | 2249 |
by uint_arith |
2250 |
||
65328 | 2251 |
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
|
2252 |
for a b c :: "'a::len word" |
65328 | 2253 |
by uint_arith |
2254 |
||
2255 |
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
|
2256 |
for x y z :: "'a::len word" |
37660 | 2257 |
by uint_arith |
2258 |
||
65328 | 2259 |
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
|
2260 |
for x y z :: "'a::len word" |
37660 | 2261 |
by uint_arith |
2262 |
||
65328 | 2263 |
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
|
2264 |
for x y z :: "'a::len word" |
37660 | 2265 |
by uint_arith |
2266 |
||
65328 | 2267 |
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
|
2268 |
for x z k :: "'a::len word" |
37660 | 2269 |
by uint_arith |
2270 |
||
65328 | 2271 |
lemma inc_le: "i < m \<Longrightarrow> i + 1 \<le> m" |
2272 |
for i m :: "'a::len word" |
|
37660 | 2273 |
by uint_arith |
2274 |
||
65328 | 2275 |
lemma inc_i: "1 \<le> i \<Longrightarrow> i < m \<Longrightarrow> 1 \<le> i + 1 \<and> i + 1 \<le> m" |
2276 |
for i m :: "'a::len word" |
|
37660 | 2277 |
by uint_arith |
2278 |
||
2279 |
lemma udvd_incr_lem: |
|
65268 | 2280 |
"up < uq \<Longrightarrow> up = ua + n * uint K \<Longrightarrow> |
65328 | 2281 |
uq = ua + n' * uint K \<Longrightarrow> up + uint K \<le> uq" |
71997 | 2282 |
by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle) |
37660 | 2283 |
|
65268 | 2284 |
lemma udvd_incr': |
2285 |
"p < q \<Longrightarrow> uint p = ua + n * uint K \<Longrightarrow> |
|
65328 | 2286 |
uint q = ua + n' * uint K \<Longrightarrow> p + K \<le> q" |
37660 | 2287 |
apply (unfold word_less_alt word_le_def) |
2288 |
apply (drule (2) udvd_incr_lem) |
|
2289 |
apply (erule uint_add_le [THEN order_trans]) |
|
2290 |
done |
|
2291 |
||
65268 | 2292 |
lemma udvd_decr': |
2293 |
"p < q \<Longrightarrow> uint p = ua + n * uint K \<Longrightarrow> |
|
65328 | 2294 |
uint q = ua + n' * uint K \<Longrightarrow> p \<le> q - K" |
37660 | 2295 |
apply (unfold word_less_alt word_le_def) |
2296 |
apply (drule (2) udvd_incr_lem) |
|
2297 |
apply (drule le_diff_eq [THEN iffD2]) |
|
2298 |
apply (erule order_trans) |
|
2299 |
apply (rule uint_sub_ge) |
|
2300 |
done |
|
2301 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
2302 |
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
|
2303 |
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
|
2304 |
lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] |
37660 | 2305 |
|
65328 | 2306 |
lemma udvd_minus_le': "xy < k \<Longrightarrow> z udvd xy \<Longrightarrow> z udvd k \<Longrightarrow> xy \<le> k - z" |
37660 | 2307 |
apply (unfold udvd_def) |
2308 |
apply clarify |
|
2309 |
apply (erule (2) udvd_decr0) |
|
2310 |
done |
|
2311 |
||
65268 | 2312 |
lemma udvd_incr2_K: |
65328 | 2313 |
"p < a + s \<Longrightarrow> a \<le> a + s \<Longrightarrow> K udvd s \<Longrightarrow> K udvd p - a \<Longrightarrow> a \<le> p \<Longrightarrow> |
2314 |
0 < K \<Longrightarrow> p \<le> p + K \<and> p + K \<le> a + s" |
|
2315 |
supply [[simproc del: linordered_ring_less_cancel_factor]] |
|
37660 | 2316 |
apply (unfold udvd_def) |
2317 |
apply clarify |
|
62390 | 2318 |
apply (simp add: uint_arith_simps split: if_split_asm) |
65268 | 2319 |
prefer 2 |
37660 | 2320 |
apply (insert uint_range' [of s])[1] |
2321 |
apply arith |
|
71997 | 2322 |
apply (drule add.commute [THEN xtrans(1)]) |
2323 |
apply (simp flip: diff_less_eq) |
|
2324 |
apply (subst (asm) mult_less_cancel_right) |
|
37660 | 2325 |
apply simp |
71997 | 2326 |
apply (simp add: diff_eq_eq not_less) |
2327 |
apply (subst (asm) (3) zless_iff_Suc_zadd) |
|
2328 |
apply auto |
|
2329 |
apply (auto simp add: algebra_simps) |
|
2330 |
apply (drule less_le_trans [of _ \<open>2 ^ LENGTH('a)\<close>]) apply assumption |
|
2331 |
apply (simp add: mult_less_0_iff) |
|
37660 | 2332 |
done |
2333 |
||
67408 | 2334 |
\<comment> \<open>links with \<open>rbl\<close> operations\<close> |
65328 | 2335 |
lemma word_succ_rbl: "to_bl w = bl \<Longrightarrow> to_bl (word_succ w) = rev (rbl_succ (rev bl))" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2336 |
apply (unfold word_succ_alt) |
37660 | 2337 |
apply clarify |
2338 |
apply (simp add: to_bl_of_bin) |
|
46654 | 2339 |
apply (simp add: to_bl_def rbl_succ) |
37660 | 2340 |
done |
2341 |
||
65328 | 2342 |
lemma word_pred_rbl: "to_bl w = bl \<Longrightarrow> to_bl (word_pred w) = rev (rbl_pred (rev bl))" |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2343 |
apply (unfold word_pred_alt) |
37660 | 2344 |
apply clarify |
2345 |
apply (simp add: to_bl_of_bin) |
|
46654 | 2346 |
apply (simp add: to_bl_def rbl_pred) |
37660 | 2347 |
done |
2348 |
||
2349 |
lemma word_add_rbl: |
|
65268 | 2350 |
"to_bl v = vbl \<Longrightarrow> to_bl w = wbl \<Longrightarrow> |
65328 | 2351 |
to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" |
37660 | 2352 |
apply (unfold word_add_def) |
2353 |
apply clarify |
|
2354 |
apply (simp add: to_bl_of_bin) |
|
2355 |
apply (simp add: to_bl_def rbl_add) |
|
2356 |
done |
|
2357 |
||
2358 |
lemma word_mult_rbl: |
|
65268 | 2359 |
"to_bl v = vbl \<Longrightarrow> to_bl w = wbl \<Longrightarrow> |
65328 | 2360 |
to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" |
37660 | 2361 |
apply (unfold word_mult_def) |
2362 |
apply clarify |
|
2363 |
apply (simp add: to_bl_of_bin) |
|
2364 |
apply (simp add: to_bl_def rbl_mult) |
|
2365 |
done |
|
2366 |
||
2367 |
lemma rtb_rbl_ariths: |
|
2368 |
"rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_succ w)) = rbl_succ ys" |
|
2369 |
"rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_pred w)) = rbl_pred ys" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2370 |
"rev (to_bl v) = ys \<Longrightarrow> rev (to_bl w) = xs \<Longrightarrow> rev (to_bl (v * w)) = rbl_mult ys xs" |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2371 |
"rev (to_bl v) = ys \<Longrightarrow> rev (to_bl w) = xs \<Longrightarrow> rev (to_bl (v + w)) = rbl_add ys xs" |
65328 | 2372 |
by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) |
37660 | 2373 |
|
2374 |
||
61799 | 2375 |
subsection \<open>Arithmetic type class instantiations\<close> |
37660 | 2376 |
|
2377 |
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
|
2378 |
word_zero_le [THEN leD, THEN antisym_conv1] |
37660 | 2379 |
|
65328 | 2380 |
lemma word_of_int_nat: "0 \<le> x \<Longrightarrow> word_of_int x = of_nat (nat x)" |
2381 |
by (simp add: word_of_int) |
|
37660 | 2382 |
|
67408 | 2383 |
text \<open> |
2384 |
note that \<open>iszero_def\<close> is only for class \<open>comm_semiring_1_cancel\<close>, |
|
2385 |
which requires word length \<open>\<ge> 1\<close>, ie \<open>'a::len word\<close> |
|
2386 |
\<close> |
|
46603 | 2387 |
lemma iszero_word_no [simp]: |
65268 | 2388 |
"iszero (numeral bin :: 'a::len word) = |
70185 | 2389 |
iszero (bintrunc (LENGTH('a)) (numeral bin))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2390 |
using word_ubin.norm_eq_iff [where 'a='a, of "numeral bin" 0] |
46603 | 2391 |
by (simp add: iszero_def [symmetric]) |
65268 | 2392 |
|
61799 | 2393 |
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
|
2394 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2395 |
lemmas word_eq_numeral_iff_iszero [simp] = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2396 |
eq_numeral_iff_iszero [where 'a="'a::len word"] |
46603 | 2397 |
|
37660 | 2398 |
|
61799 | 2399 |
subsection \<open>Word and nat\<close> |
37660 | 2400 |
|
45811 | 2401 |
lemma td_ext_unat [OF refl]: |
70185 | 2402 |
"n = LENGTH('a::len) \<Longrightarrow> |
65328 | 2403 |
td_ext (unat :: 'a word \<Rightarrow> nat) of_nat (unats n) (\<lambda>i. i mod 2 ^ n)" |
37660 | 2404 |
apply (unfold td_ext_def' unat_def word_of_nat unats_uints) |
2405 |
apply (auto intro!: imageI simp add : word_of_int_hom_syms) |
|
65328 | 2406 |
apply (erule word_uint.Abs_inverse [THEN arg_cong]) |
37660 | 2407 |
apply (simp add: int_word_uint nat_mod_distrib nat_power_eq) |
2408 |
done |
|
2409 |
||
45604 | 2410 |
lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] |
37660 | 2411 |
|
2412 |
interpretation word_unat: |
|
65328 | 2413 |
td_ext |
2414 |
"unat::'a::len word \<Rightarrow> nat" |
|
2415 |
of_nat |
|
70185 | 2416 |
"unats (LENGTH('a::len))" |
2417 |
"\<lambda>i. i mod 2 ^ LENGTH('a::len)" |
|
37660 | 2418 |
by (rule td_ext_unat) |
2419 |
||
2420 |
lemmas td_unat = word_unat.td_thm |
|
2421 |
||
2422 |
lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] |
|
2423 |
||
70185 | 2424 |
lemma unat_le: "y \<le> unat z \<Longrightarrow> y \<in> unats (LENGTH('a))" |
65328 | 2425 |
for z :: "'a::len word" |
37660 | 2426 |
apply (unfold unats_def) |
2427 |
apply clarsimp |
|
65268 | 2428 |
apply (rule xtrans, rule unat_lt2p, assumption) |
37660 | 2429 |
done |
2430 |
||
70185 | 2431 |
lemma word_nchotomy: "\<forall>w :: 'a::len word. \<exists>n. w = of_nat n \<and> n < 2 ^ LENGTH('a)" |
37660 | 2432 |
apply (rule allI) |
2433 |
apply (rule word_unat.Abs_cases) |
|
2434 |
apply (unfold unats_def) |
|
2435 |
apply auto |
|
2436 |
done |
|
2437 |
||
70185 | 2438 |
lemma of_nat_eq: "of_nat n = w \<longleftrightarrow> (\<exists>q. n = unat w + q * 2 ^ LENGTH('a))" |
65328 | 2439 |
for w :: "'a::len word" |
68157 | 2440 |
using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] |
2441 |
by (auto simp add: word_unat.inverse_norm) |
|
37660 | 2442 |
|
65328 | 2443 |
lemma of_nat_eq_size: "of_nat n = w \<longleftrightarrow> (\<exists>q. n = unat w + q * 2 ^ size w)" |
37660 | 2444 |
unfolding word_size by (rule of_nat_eq) |
2445 |
||
70185 | 2446 |
lemma of_nat_0: "of_nat m = (0::'a::len word) \<longleftrightarrow> (\<exists>q. m = q * 2 ^ LENGTH('a))" |
37660 | 2447 |
by (simp add: of_nat_eq) |
2448 |
||
70185 | 2449 |
lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" |
45805 | 2450 |
by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) |
37660 | 2451 |
|
65328 | 2452 |
lemma of_nat_gt_0: "of_nat k \<noteq> 0 \<Longrightarrow> 0 < k" |
37660 | 2453 |
by (cases k) auto |
2454 |
||
70185 | 2455 |
lemma of_nat_neq_0: "0 < k \<Longrightarrow> k < 2 ^ LENGTH('a::len) \<Longrightarrow> of_nat k \<noteq> (0 :: 'a word)" |
65328 | 2456 |
by (auto simp add : of_nat_0) |
2457 |
||
2458 |
lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" |
|
37660 | 2459 |
by simp |
2460 |
||
65328 | 2461 |
lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
2462 |
by (simp add: word_of_nat wi_hom_mult) |
37660 | 2463 |
|
65328 | 2464 |
lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
2465 |
by (simp add: word_of_nat wi_hom_succ ac_simps) |
37660 | 2466 |
|
2467 |
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
|
2468 |
by simp |
37660 | 2469 |
|
2470 |
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
|
2471 |
by simp |
37660 | 2472 |
|
65268 | 2473 |
lemmas Abs_fnat_homs = |
2474 |
Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc |
|
37660 | 2475 |
Abs_fnat_hom_0 Abs_fnat_hom_1 |
2476 |
||
65328 | 2477 |
lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" |
2478 |
by simp |
|
2479 |
||
2480 |
lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" |
|
37660 | 2481 |
by simp |
2482 |
||
65328 | 2483 |
lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" |
37660 | 2484 |
by (subst Abs_fnat_hom_Suc [symmetric]) simp |
2485 |
||
65328 | 2486 |
lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" |
37660 | 2487 |
by (simp add: word_div_def word_of_nat zdiv_int uint_nat) |
2488 |
||
65328 | 2489 |
lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" |
37660 | 2490 |
by (simp add: word_mod_def word_of_nat zmod_int uint_nat) |
2491 |
||
2492 |
lemmas word_arith_nat_defs = |
|
2493 |
word_arith_nat_add word_arith_nat_mult |
|
2494 |
word_arith_nat_Suc Abs_fnat_hom_0 |
|
2495 |
Abs_fnat_hom_1 word_arith_nat_div |
|
65268 | 2496 |
word_arith_nat_mod |
37660 | 2497 |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
2498 |
lemma unat_cong: "x = y \<Longrightarrow> unat x = unat y" |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
2499 |
by simp |
65268 | 2500 |
|
37660 | 2501 |
lemmas unat_word_ariths = word_arith_nat_defs |
45604 | 2502 |
[THEN trans [OF unat_cong unat_of_nat]] |
37660 | 2503 |
|
2504 |
lemmas word_sub_less_iff = word_sub_le_iff |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
2505 |
[unfolded linorder_not_less [symmetric] Not_eq_iff] |
37660 | 2506 |
|
65268 | 2507 |
lemma unat_add_lem: |
70185 | 2508 |
"unat x + unat y < 2 ^ LENGTH('a) \<longleftrightarrow> unat (x + y) = unat x + unat y" |
65328 | 2509 |
for x y :: "'a::len word" |
71997 | 2510 |
apply (auto simp: unat_word_ariths) |
2511 |
apply (metis unat_lt2p word_unat.eq_norm) |
|
2512 |
done |
|
37660 | 2513 |
|
65268 | 2514 |
lemma unat_mult_lem: |
70185 | 2515 |
"unat x * unat y < 2 ^ LENGTH('a) \<longleftrightarrow> unat (x * y) = unat x * unat y" |
65363 | 2516 |
for x y :: "'a::len word" |
71997 | 2517 |
apply (auto simp: unat_word_ariths) |
2518 |
apply (metis unat_lt2p word_unat.eq_norm) |
|
2519 |
done |
|
2520 |
||
2521 |
lemma unat_plus_if': |
|
2522 |
\<open>unat (a + b) = |
|
2523 |
(if unat a + unat b < 2 ^ LENGTH('a) |
|
2524 |
then unat a + unat b |
|
2525 |
else unat a + unat b - 2 ^ LENGTH('a))\<close> for a b :: \<open>'a::len word\<close> |
|
2526 |
apply (auto simp: unat_word_ariths not_less) |
|
2527 |
apply (subst (asm) le_iff_add) |
|
2528 |
apply auto |
|
2529 |
apply (metis add_less_cancel_left add_less_cancel_right le_less_trans less_imp_le mod_less unat_lt2p) |
|
2530 |
done |
|
65328 | 2531 |
|
2532 |
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
|
2533 |
for a b x :: "'a::len word" |
37660 | 2534 |
apply (erule order_trans) |
2535 |
apply (erule olen_add_eqv [THEN iffD1]) |
|
2536 |
done |
|
2537 |
||
65328 | 2538 |
lemmas un_ui_le = |
2539 |
trans [OF word_le_nat_alt [symmetric] word_le_def] |
|
37660 | 2540 |
|
2541 |
lemma unat_sub_if_size: |
|
65328 | 2542 |
"unat (x - y) = |
2543 |
(if unat y \<le> unat x |
|
2544 |
then unat x - unat y |
|
2545 |
else unat x + 2 ^ size x - unat y)" |
|
37660 | 2546 |
apply (unfold word_size) |
2547 |
apply (simp add: un_ui_le) |
|
2548 |
apply (auto simp add: unat_def uint_sub_if') |
|
2549 |
apply (rule nat_diff_distrib) |
|
2550 |
prefer 3 |
|
2551 |
apply (simp add: algebra_simps) |
|
2552 |
apply (rule nat_diff_distrib [THEN trans]) |
|
2553 |
prefer 3 |
|
2554 |
apply (subst nat_add_distrib) |
|
2555 |
prefer 3 |
|
2556 |
apply (simp add: nat_power_eq) |
|
2557 |
apply auto |
|
2558 |
apply uint_arith |
|
2559 |
done |
|
2560 |
||
2561 |
lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] |
|
2562 |
||
71997 | 2563 |
lemma uint_div: |
2564 |
\<open>uint (x div y) = uint x div uint y\<close> |
|
2565 |
by (metis div_le_dividend le_less_trans mod_less uint_nat unat_lt2p unat_word_ariths(6) zdiv_int) |
|
2566 |
||
2567 |
lemma unat_div: |
|
2568 |
\<open>unat (x div y) = unat x div unat y\<close> |
|
2569 |
by (simp add: unat_def uint_div add: nat_div_distrib) |
|
2570 |
||
2571 |
lemma uint_mod: |
|
2572 |
\<open>uint (x mod y) = uint x mod uint y\<close> |
|
2573 |
by (metis (no_types, lifting) le_less_trans mod_by_0 mod_le_divisor mod_less neq0_conv uint_nat unat_lt2p unat_word_ariths(7) zmod_int) |
|
2574 |
||
65328 | 2575 |
lemma unat_mod: "unat (x mod y) = unat x mod unat y" |
2576 |
for x y :: "'a::len word" |
|
71997 | 2577 |
by (simp add: unat_def uint_mod add: nat_mod_distrib) |
2578 |
||
37660 | 2579 |
|
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
2580 |
text \<open>Definition of \<open>unat_arith\<close> tactic\<close> |
37660 | 2581 |
|
70185 | 2582 |
lemma unat_split: "P (unat x) \<longleftrightarrow> (\<forall>n. of_nat n = x \<and> n < 2^LENGTH('a) \<longrightarrow> P n)" |
65328 | 2583 |
for x :: "'a::len word" |
37660 | 2584 |
by (auto simp: unat_of_nat) |
2585 |
||
70185 | 2586 |
lemma unat_split_asm: "P (unat x) \<longleftrightarrow> (\<nexists>n. of_nat n = x \<and> n < 2^LENGTH('a) \<and> \<not> P n)" |
65328 | 2587 |
for x :: "'a::len word" |
37660 | 2588 |
by (auto simp: unat_of_nat) |
2589 |
||
65268 | 2590 |
lemmas of_nat_inverse = |
37660 | 2591 |
word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] |
2592 |
||
2593 |
lemmas unat_splits = unat_split unat_split_asm |
|
2594 |
||
2595 |
lemmas unat_arith_simps = |
|
2596 |
word_le_nat_alt word_less_nat_alt |
|
2597 |
word_unat.Rep_inject [symmetric] |
|
2598 |
unat_sub_if' unat_plus_if' unat_div unat_mod |
|
2599 |
||
67408 | 2600 |
\<comment> \<open>\<open>unat_arith_tac\<close>: tactic to reduce word arithmetic to \<open>nat\<close>, try to solve via \<open>arith\<close>\<close> |
61799 | 2601 |
ML \<open> |
65268 | 2602 |
fun unat_arith_simpset ctxt = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2603 |
ctxt addsimps @{thms unat_arith_simps} |
37660 | 2604 |
delsimps @{thms word_unat.Rep_inject} |
62390 | 2605 |
|> fold Splitter.add_split @{thms if_split_asm} |
45620
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45604
diff
changeset
|
2606 |
|> fold Simplifier.add_cong @{thms power_False_cong} |
37660 | 2607 |
|
65268 | 2608 |
fun unat_arith_tacs ctxt = |
37660 | 2609 |
let |
2610 |
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
|
2611 |
Arith_Data.arith_tac ctxt n t |
37660 | 2612 |
handle Cooper.COOPER _ => Seq.empty; |
65268 | 2613 |
in |
42793 | 2614 |
[ clarify_tac ctxt 1, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2615 |
full_simp_tac (unat_arith_simpset ctxt) 1, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2616 |
ALLGOALS (full_simp_tac |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2617 |
(put_simpset HOL_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2618 |
|> fold Splitter.add_split @{thms unat_splits} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51375
diff
changeset
|
2619 |
|> fold Simplifier.add_cong @{thms power_False_cong})), |
65268 | 2620 |
rewrite_goals_tac ctxt @{thms word_size}, |
60754 | 2621 |
ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN |
2622 |
REPEAT (eresolve_tac ctxt [conjE] n) THEN |
|
2623 |
REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), |
|
65268 | 2624 |
TRYALL arith_tac' ] |
37660 | 2625 |
end |
2626 |
||
2627 |
fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) |
|
61799 | 2628 |
\<close> |
37660 | 2629 |
|
65268 | 2630 |
method_setup unat_arith = |
61799 | 2631 |
\<open>Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\<close> |
37660 | 2632 |
"solving word arithmetic via natural numbers and arith" |
2633 |
||
65328 | 2634 |
lemma no_plus_overflow_unat_size: "x \<le> x + y \<longleftrightarrow> unat x + unat y < 2 ^ size x" |
2635 |
for x y :: "'a::len word" |
|
37660 | 2636 |
unfolding word_size by unat_arith |
2637 |
||
65328 | 2638 |
lemmas no_olen_add_nat = |
2639 |
no_plus_overflow_unat_size [unfolded word_size] |
|
2640 |
||
2641 |
lemmas unat_plus_simple = |
|
2642 |
trans [OF no_olen_add_nat unat_add_lem] |
|
2643 |
||
70185 | 2644 |
lemma word_div_mult: "0 < y \<Longrightarrow> unat x * unat y < 2 ^ LENGTH('a) \<Longrightarrow> x * y div y = x" |
65328 | 2645 |
for x y :: "'a::len word" |
37660 | 2646 |
apply unat_arith |
2647 |
apply clarsimp |
|
2648 |
apply (subst unat_mult_lem [THEN iffD1]) |
|
65328 | 2649 |
apply auto |
37660 | 2650 |
done |
2651 |
||
70185 | 2652 |
lemma div_lt': "i \<le> k div x \<Longrightarrow> unat i * unat x < 2 ^ LENGTH('a)" |
65328 | 2653 |
for i k x :: "'a::len word" |
37660 | 2654 |
apply unat_arith |
2655 |
apply clarsimp |
|
2656 |
apply (drule mult_le_mono1) |
|
2657 |
apply (erule order_le_less_trans) |
|
71997 | 2658 |
apply (metis add_lessD1 div_mult_mod_eq unat_lt2p) |
37660 | 2659 |
done |
2660 |
||
2661 |
lemmas div_lt'' = order_less_imp_le [THEN div_lt'] |
|
2662 |
||
65328 | 2663 |
lemma div_lt_mult: "i < k div x \<Longrightarrow> 0 < x \<Longrightarrow> i * x < k" |
2664 |
for i k x :: "'a::len word" |
|
37660 | 2665 |
apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) |
2666 |
apply (simp add: unat_arith_simps) |
|
2667 |
apply (drule (1) mult_less_mono1) |
|
2668 |
apply (erule order_less_le_trans) |
|
71997 | 2669 |
apply auto |
37660 | 2670 |
done |
2671 |
||
65328 | 2672 |
lemma div_le_mult: "i \<le> k div x \<Longrightarrow> 0 < x \<Longrightarrow> i * x \<le> k" |
2673 |
for i k x :: "'a::len word" |
|
37660 | 2674 |
apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) |
2675 |
apply (simp add: unat_arith_simps) |
|
2676 |
apply (drule mult_le_mono1) |
|
2677 |
apply (erule order_trans) |
|
71997 | 2678 |
apply auto |
37660 | 2679 |
done |
2680 |
||
70185 | 2681 |
lemma div_lt_uint': "i \<le> k div x \<Longrightarrow> uint i * uint x < 2 ^ LENGTH('a)" |
65328 | 2682 |
for i k x :: "'a::len word" |
37660 | 2683 |
apply (unfold uint_nat) |
2684 |
apply (drule div_lt') |
|
65328 | 2685 |
apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) |
2686 |
done |
|
37660 | 2687 |
|
2688 |
lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] |
|
2689 |
||
70185 | 2690 |
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
|
2691 |
for x y z :: "'a::len word" |
71997 | 2692 |
by (metis add_diff_cancel_left' add_diff_eq uint_add_lem uint_plus_simple) |
2693 |
||
37660 | 2694 |
lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] |
2695 |
||
2696 |
lemmas plus_minus_no_overflow = |
|
2697 |
order_less_imp_le [THEN plus_minus_no_overflow_ab] |
|
65268 | 2698 |
|
37660 | 2699 |
lemmas mcs = word_less_minus_cancel word_less_minus_mono_left |
2700 |
word_le_minus_cancel word_le_minus_mono_left |
|
2701 |
||
45604 | 2702 |
lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x |
2703 |
lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x |
|
2704 |
lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x |
|
37660 | 2705 |
|
2706 |
lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] |
|
2707 |
||
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66453
diff
changeset
|
2708 |
lemmas thd = times_div_less_eq_dividend |
37660 | 2709 |
|
71997 | 2710 |
lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend |
37660 | 2711 |
|
65328 | 2712 |
lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" |
2713 |
for n b :: "'a::len word" |
|
71997 | 2714 |
by (fact div_mult_mod_eq) |
37660 | 2715 |
|
65328 | 2716 |
lemma word_div_mult_le: "a div b * b \<le> a" |
2717 |
for a b :: "'a::len word" |
|
71997 | 2718 |
by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) |
37660 | 2719 |
|
65328 | 2720 |
lemma word_mod_less_divisor: "0 < n \<Longrightarrow> m mod n < n" |
2721 |
for m n :: "'a::len word" |
|
71997 | 2722 |
by (simp add: unat_arith_simps) |
2723 |
||
65328 | 2724 |
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
|
2725 |
by (induct n) (simp_all add: wi_hom_mult [symmetric]) |
37660 | 2726 |
|
65328 | 2727 |
lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" |
37660 | 2728 |
by (simp add : word_of_int_power_hom [symmetric]) |
2729 |
||
65268 | 2730 |
lemma of_bl_length_less: |
70185 | 2731 |
"length x = k \<Longrightarrow> k < LENGTH('a) \<Longrightarrow> (of_bl x :: 'a::len word) < 2 ^ k" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2732 |
apply (unfold of_bl_def word_less_alt word_numeral_alt) |
37660 | 2733 |
apply safe |
65268 | 2734 |
apply (simp (no_asm) add: word_of_int_power_hom word_uint.eq_norm |
65328 | 2735 |
del: word_of_int_numeral) |
71942 | 2736 |
apply simp |
37660 | 2737 |
apply (subst mod_pos_pos_trivial) |
2738 |
apply (rule bl_to_bin_ge0) |
|
2739 |
apply (rule order_less_trans) |
|
2740 |
apply (rule bl_to_bin_lt2p) |
|
2741 |
apply simp |
|
46646 | 2742 |
apply (rule bl_to_bin_lt2p) |
37660 | 2743 |
done |
2744 |
||
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
2745 |
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
|
2746 |
for n :: "'a::len word" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
2747 |
by unat_arith |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
2748 |
|
37660 | 2749 |
|
61799 | 2750 |
subsection \<open>Cardinality, finiteness of set of words\<close> |
37660 | 2751 |
|
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
|
2752 |
lemma inj_on_word_of_int: \<open>inj_on (word_of_int :: int \<Rightarrow> 'a word) {0..<2 ^ LENGTH('a::len)}\<close> |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2753 |
by (rule inj_onI) (simp add: word.abs_eq_iff take_bit_eq_mod) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2754 |
|
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2755 |
lemma inj_uint: \<open>inj uint\<close> |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2756 |
by (rule injI) simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2757 |
|
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
|
2758 |
lemma range_uint: \<open>range (uint :: 'a word \<Rightarrow> int) = {0..<2 ^ LENGTH('a::len)}\<close> |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2759 |
by transfer (auto simp add: bintr_lt2p range_bintrunc) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2760 |
|
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
|
2761 |
lemma UNIV_eq: \<open>(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\<close> |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2762 |
proof - |
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
|
2763 |
have \<open>uint ` (UNIV :: 'a word set) = uint ` (word_of_int :: int \<Rightarrow> 'a word) ` {0..<2 ^ LENGTH('a::len)}\<close> |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2764 |
by (simp add: range_uint image_image uint.abs_eq take_bit_eq_mod) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2765 |
then show ?thesis |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2766 |
using inj_image_eq_iff [of \<open>uint :: 'a word \<Rightarrow> int\<close> \<open>UNIV :: 'a word set\<close> \<open>word_of_int ` {0..<2 ^ LENGTH('a)} :: 'a word set\<close>, OF inj_uint] |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2767 |
by simp |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2768 |
qed |
45809
2bee94cbae72
finite class instance for word type; remove unused lemmas
huffman
parents:
45808
diff
changeset
|
2769 |
|
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
|
2770 |
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
|
2771 |
by (simp add: UNIV_eq card_image inj_on_word_of_int) |
37660 | 2772 |
|
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
2773 |
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
|
2774 |
for x :: "'a::len word" |
65328 | 2775 |
unfolding word_size by (rule card_word) |
37660 | 2776 |
|
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
|
2777 |
instance word :: (len) finite |
71948
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2778 |
by standard (simp add: UNIV_eq) |
6ede899d26d3
fundamental construction of word type following existing transfer rules
haftmann
parents:
71947
diff
changeset
|
2779 |
|
37660 | 2780 |
|
61799 | 2781 |
subsection \<open>Bitwise Operations on Words\<close> |
37660 | 2782 |
|
70190 | 2783 |
lemma word_eq_rbl_eq: "x = y \<longleftrightarrow> rev (to_bl x) = rev (to_bl y)" |
2784 |
by simp |
|
2785 |
||
37660 | 2786 |
lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or |
65268 | 2787 |
|
67408 | 2788 |
\<comment> \<open>following definitions require both arithmetic and bit-wise word operations\<close> |
2789 |
||
2790 |
\<comment> \<open>to get \<open>word_no_log_defs\<close> from \<open>word_log_defs\<close>, using \<open>bin_log_bintrs\<close>\<close> |
|
37660 | 2791 |
lemmas wils1 = bin_log_bintrs [THEN word_ubin.norm_eq_iff [THEN iffD1], |
45604 | 2792 |
folded word_ubin.eq_norm, THEN eq_reflection] |
37660 | 2793 |
|
67408 | 2794 |
\<comment> \<open>the binary operations only\<close> (* BH: why is this needed? *) |
65268 | 2795 |
lemmas word_log_binary_defs = |
37660 | 2796 |
word_and_def word_or_def word_xor_def |
2797 |
||
46011 | 2798 |
lemma word_wi_log_defs: |
71149 | 2799 |
"NOT (word_of_int a) = word_of_int (NOT a)" |
46011 | 2800 |
"word_of_int a AND word_of_int b = word_of_int (a AND b)" |
2801 |
"word_of_int a OR word_of_int b = word_of_int (a OR b)" |
|
2802 |
"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
|
2803 |
by (transfer, rule refl)+ |
47372 | 2804 |
|
46011 | 2805 |
lemma word_no_log_defs [simp]: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2806 |
"NOT (numeral a) = word_of_int (NOT (numeral a))" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
2807 |
"NOT (- numeral a) = word_of_int (NOT (- numeral a))" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2808 |
"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
|
2809 |
"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
|
2810 |
"- 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
|
2811 |
"- 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
|
2812 |
"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
|
2813 |
"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
|
2814 |
"- 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
|
2815 |
"- 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
|
2816 |
"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
|
2817 |
"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
|
2818 |
"- 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
|
2819 |
"- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" |
47372 | 2820 |
by (transfer, rule refl)+ |
37660 | 2821 |
|
61799 | 2822 |
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
|
2823 |
|
88ef116e0522
add simp rules for bitwise word operations with 1
huffman
parents:
46057
diff
changeset
|
2824 |
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
|
2825 |
"NOT (1::'a::len word) = -2" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2826 |
"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
|
2827 |
"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
|
2828 |
"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
|
2829 |
"- 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
|
2830 |
"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
|
2831 |
"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
|
2832 |
"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
|
2833 |
"- 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
|
2834 |
"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
|
2835 |
"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
|
2836 |
"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
|
2837 |
"- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" |
47372 | 2838 |
by (transfer, simp)+ |
46064
88ef116e0522
add simp rules for bitwise word operations with 1
huffman
parents:
46057
diff
changeset
|
2839 |
|
61799 | 2840 |
text \<open>Special cases for when one of the arguments equals -1.\<close> |
56979 | 2841 |
|
2842 |
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
|
2843 |
"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
|
2844 |
"(-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
|
2845 |
"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
|
2846 |
"(-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
|
2847 |
"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
|
2848 |
" (-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
|
2849 |
"x XOR (-1::'a::len word) = NOT x" |
56979 | 2850 |
by (transfer, simp)+ |
2851 |
||
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2852 |
lemma uint_and: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2853 |
\<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
|
2854 |
by transfer simp |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2855 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2856 |
lemma uint_or: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2857 |
\<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
|
2858 |
by transfer simp |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2859 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2860 |
lemma uint_xor: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2861 |
\<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
|
2862 |
by transfer simp |
47372 | 2863 |
|
2864 |
lemma test_bit_wi [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
|
2865 |
"(word_of_int x :: 'a::len word) !! n \<longleftrightarrow> n < LENGTH('a) \<and> bin_nth x n" |
65328 | 2866 |
by (simp add: word_test_bit_def word_ubin.eq_norm nth_bintr) |
47372 | 2867 |
|
2868 |
lemma word_test_bit_transfer [transfer_rule]: |
|
67399 | 2869 |
"(rel_fun pcr_word (rel_fun (=) (=))) |
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
|
2870 |
(\<lambda>x n. n < LENGTH('a) \<and> bin_nth x n) (test_bit :: 'a::len word \<Rightarrow> _)" |
55945 | 2871 |
unfolding rel_fun_def word.pcr_cr_eq cr_word_def by simp |
37660 | 2872 |
|
2873 |
lemma word_ops_nth_size: |
|
65328 | 2874 |
"n < size x \<Longrightarrow> |
2875 |
(x OR y) !! n = (x !! n | y !! n) \<and> |
|
2876 |
(x AND y) !! n = (x !! n \<and> y !! n) \<and> |
|
2877 |
(x XOR y) !! n = (x !! n \<noteq> y !! n) \<and> |
|
2878 |
(NOT x) !! n = (\<not> x !! 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
|
2879 |
for x :: "'a::len word" |
47372 | 2880 |
unfolding word_size by transfer (simp add: bin_nth_ops) |
37660 | 2881 |
|
2882 |
lemma word_ao_nth: |
|
65328 | 2883 |
"(x OR y) !! n = (x !! n | y !! n) \<and> |
2884 |
(x AND y) !! n = (x !! n \<and> y !! 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
|
2885 |
for x :: "'a::len word" |
47372 | 2886 |
by transfer (auto simp add: bin_nth_ops) |
46023
fad87bb608fc
restate some lemmas to respect int/bin distinction
huffman
parents:
46022
diff
changeset
|
2887 |
|
72000 | 2888 |
lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] |
2889 |
lemmas msb1 = msb0 [where i = 0] |
|
2890 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2891 |
lemma test_bit_numeral [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
|
2892 |
"(numeral w :: 'a::len word) !! n \<longleftrightarrow> |
70185 | 2893 |
n < LENGTH('a) \<and> bin_nth (numeral w) n" |
47372 | 2894 |
by transfer (rule refl) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2895 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2896 |
lemma test_bit_neg_numeral [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
|
2897 |
"(- numeral w :: 'a::len word) !! n \<longleftrightarrow> |
70185 | 2898 |
n < LENGTH('a) \<and> bin_nth (- numeral w) n" |
47372 | 2899 |
by transfer (rule refl) |
46023
fad87bb608fc
restate some lemmas to respect int/bin distinction
huffman
parents:
46022
diff
changeset
|
2900 |
|
65328 | 2901 |
lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \<longleftrightarrow> n = 0" |
47372 | 2902 |
by transfer auto |
65268 | 2903 |
|
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
|
2904 |
lemma nth_0 [simp]: "\<not> (0 :: 'a::len word) !! n" |
47372 | 2905 |
by transfer simp |
46023
fad87bb608fc
restate some lemmas to respect int/bin distinction
huffman
parents:
46022
diff
changeset
|
2906 |
|
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
|
2907 |
lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \<longleftrightarrow> n < LENGTH('a)" |
47372 | 2908 |
by transfer simp |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
2909 |
|
67408 | 2910 |
\<comment> \<open>get from commutativity, associativity etc of \<open>int_and\<close> etc to same for \<open>word_and etc\<close>\<close> |
65268 | 2911 |
lemmas bwsimps = |
46013 | 2912 |
wi_hom_add |
37660 | 2913 |
word_wi_log_defs |
2914 |
||
2915 |
lemma word_bw_assocs: |
|
2916 |
"(x AND y) AND z = x AND y AND z" |
|
2917 |
"(x OR y) OR z = x OR y OR z" |
|
2918 |
"(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
|
2919 |
for x :: "'a::len word" |
46022 | 2920 |
by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) |
65268 | 2921 |
|
37660 | 2922 |
lemma word_bw_comms: |
2923 |
"x AND y = y AND x" |
|
2924 |
"x OR y = y OR x" |
|
2925 |
"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
|
2926 |
for x :: "'a::len word" |
46022 | 2927 |
by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) |
65268 | 2928 |
|
37660 | 2929 |
lemma word_bw_lcs: |
2930 |
"y AND x AND z = x AND y AND z" |
|
2931 |
"y OR x OR z = x OR y OR z" |
|
2932 |
"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
|
2933 |
for x :: "'a::len word" |
46022 | 2934 |
by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) |
37660 | 2935 |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2936 |
lemma word_log_esimps: |
37660 | 2937 |
"x AND 0 = 0" |
2938 |
"x AND -1 = x" |
|
2939 |
"x OR 0 = x" |
|
2940 |
"x OR -1 = -1" |
|
2941 |
"x XOR 0 = x" |
|
2942 |
"x XOR -1 = NOT x" |
|
2943 |
"0 AND x = 0" |
|
2944 |
"-1 AND x = x" |
|
2945 |
"0 OR x = x" |
|
2946 |
"-1 OR x = -1" |
|
2947 |
"0 XOR x = x" |
|
2948 |
"-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
|
2949 |
for x :: "'a::len word" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2950 |
by simp_all |
37660 | 2951 |
|
2952 |
lemma word_not_dist: |
|
2953 |
"NOT (x OR y) = NOT x AND NOT y" |
|
2954 |
"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
|
2955 |
for x :: "'a::len word" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2956 |
by simp_all |
37660 | 2957 |
|
2958 |
lemma word_bw_same: |
|
2959 |
"x AND x = x" |
|
2960 |
"x OR x = x" |
|
2961 |
"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
|
2962 |
for x :: "'a::len word" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2963 |
by simp_all |
37660 | 2964 |
|
2965 |
lemma word_ao_absorbs [simp]: |
|
2966 |
"x AND (y OR x) = x" |
|
2967 |
"x OR y AND x = x" |
|
2968 |
"x AND (x OR y) = x" |
|
2969 |
"y AND x OR x = x" |
|
2970 |
"(y OR x) AND x = x" |
|
2971 |
"x OR x AND y = x" |
|
2972 |
"(x OR y) AND x = x" |
|
2973 |
"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
|
2974 |
for x :: "'a::len word" |
46022 | 2975 |
by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) |
37660 | 2976 |
|
71149 | 2977 |
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
|
2978 |
for x :: "'a::len word" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
2979 |
by simp |
37660 | 2980 |
|
65328 | 2981 |
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
|
2982 |
for x :: "'a::len word" |
46022 | 2983 |
by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) |
37660 | 2984 |
|
65328 | 2985 |
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
|
2986 |
for x :: "'a::len word" |
65328 | 2987 |
by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) |
2988 |
||
2989 |
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
|
2990 |
for x :: "'a::len word" |
47372 | 2991 |
by transfer (simp add: bin_add_not) |
37660 | 2992 |
|
65328 | 2993 |
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
|
2994 |
for x :: "'a::len word" |
47372 | 2995 |
by transfer (simp add: plus_and_or) |
37660 | 2996 |
|
65328 | 2997 |
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
|
2998 |
for x :: "'a::len word" |
65328 | 2999 |
by auto |
3000 |
||
3001 |
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
|
3002 |
for x' :: "'a::len word" |
65328 | 3003 |
by auto |
3004 |
||
3005 |
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
|
3006 |
for w w' :: "'a::len word" |
48196 | 3007 |
by (auto intro: leoa leao) |
37660 | 3008 |
|
65328 | 3009 |
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
|
3010 |
for x y :: "'a::len word" |
65328 | 3011 |
by (auto simp: word_le_def uint_or intro: le_int_or) |
37660 | 3012 |
|
71997 | 3013 |
lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] |
3014 |
lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] |
|
3015 |
lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] |
|
37660 | 3016 |
|
65268 | 3017 |
lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" |
45550
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3018 |
unfolding to_bl_def word_log_defs bl_not_bin |
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3019 |
by (simp add: word_ubin.eq_norm) |
37660 | 3020 |
|
67399 | 3021 |
lemma bl_word_xor: "to_bl (v XOR w) = map2 (\<noteq>) (to_bl v) (to_bl w)" |
37660 | 3022 |
unfolding to_bl_def word_log_defs bl_xor_bin |
45550
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3023 |
by (simp add: word_ubin.eq_norm) |
37660 | 3024 |
|
67399 | 3025 |
lemma bl_word_or: "to_bl (v OR w) = map2 (\<or>) (to_bl v) (to_bl w)" |
45550
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3026 |
unfolding to_bl_def word_log_defs bl_or_bin |
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3027 |
by (simp add: word_ubin.eq_norm) |
37660 | 3028 |
|
67399 | 3029 |
lemma bl_word_and: "to_bl (v AND w) = map2 (\<and>) (to_bl v) (to_bl w)" |
45550
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3030 |
unfolding to_bl_def word_log_defs bl_and_bin |
73a4f31d41c4
Word.thy: reduce usage of numeral-representation-dependent thms like number_of_is_id in proofs
huffman
parents:
45549
diff
changeset
|
3031 |
by (simp add: word_ubin.eq_norm) |
37660 | 3032 |
|
65328 | 3033 |
lemma bin_nth_uint': "bin_nth (uint w) n \<longleftrightarrow> rev (bin_to_bl (size w) (uint w)) ! n \<and> n < size w" |
37660 | 3034 |
apply (unfold word_size) |
3035 |
apply (safe elim!: bin_nth_uint_imp) |
|
3036 |
apply (frule bin_nth_uint_imp) |
|
3037 |
apply (fast dest!: bin_nth_bl)+ |
|
3038 |
done |
|
3039 |
||
3040 |
lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] |
|
3041 |
||
65328 | 3042 |
lemma test_bit_bl: "w !! n \<longleftrightarrow> rev (to_bl w) ! n \<and> n < size w" |
3043 |
unfolding to_bl_def word_test_bit_def word_size by (rule bin_nth_uint) |
|
37660 | 3044 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3045 |
lemma to_bl_nth: "n < size w \<Longrightarrow> to_bl w ! n = w !! (size w - Suc n)" |
71997 | 3046 |
by (simp add: word_size rev_nth test_bit_bl) |
37660 | 3047 |
|
71990 | 3048 |
lemma map_bit_interval_eq: |
3049 |
\<open>map (bit w) [0..<n] = takefill False n (rev (to_bl w))\<close> for w :: \<open>'a::len word\<close> |
|
3050 |
proof (rule nth_equalityI) |
|
3051 |
show \<open>length (map (bit w) [0..<n]) = |
|
3052 |
length (takefill False n (rev (to_bl w)))\<close> |
|
3053 |
by simp |
|
3054 |
fix m |
|
3055 |
assume \<open>m < length (map (bit w) [0..<n])\<close> |
|
3056 |
then have \<open>m < n\<close> |
|
3057 |
by simp |
|
3058 |
then have \<open>bit w m \<longleftrightarrow> takefill False n (rev (to_bl w)) ! m\<close> |
|
3059 |
by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length) |
|
3060 |
with \<open>m < n \<close>show \<open>map (bit w) [0..<n] ! m \<longleftrightarrow> takefill False n (rev (to_bl w)) ! m\<close> |
|
3061 |
by simp |
|
3062 |
qed |
|
3063 |
||
3064 |
lemma to_bl_unfold: |
|
3065 |
\<open>to_bl w = rev (map (bit w) [0..<LENGTH('a)])\<close> for w :: \<open>'a::len word\<close> |
|
3066 |
by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) |
|
3067 |
||
3068 |
lemma nth_rev_to_bl: |
|
3069 |
\<open>rev (to_bl w) ! n \<longleftrightarrow> bit w n\<close> |
|
3070 |
if \<open>n < LENGTH('a)\<close> for w :: \<open>'a::len word\<close> |
|
3071 |
using that by (simp add: to_bl_unfold) |
|
3072 |
||
3073 |
lemma nth_to_bl: |
|
3074 |
\<open>to_bl w ! n \<longleftrightarrow> bit w (LENGTH('a) - Suc n)\<close> |
|
3075 |
if \<open>n < LENGTH('a)\<close> for w :: \<open>'a::len word\<close> |
|
3076 |
using that by (simp add: to_bl_unfold rev_nth) |
|
3077 |
||
37660 | 3078 |
lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" |
65363 | 3079 |
by (auto simp: of_bl_def bl_to_bin_rep_F) |
65268 | 3080 |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3081 |
lemma bit_horner_sum_bit_word_iff: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3082 |
\<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
|
3083 |
\<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
|
3084 |
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
|
3085 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3086 |
lemma of_bl_eq: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3087 |
\<open>of_bl bs = horner_sum of_bool 2 (rev bs)\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3088 |
by (rule bit_word_eqI) (simp add: bit_of_bl_iff bit_horner_sum_bit_word_iff ac_simps) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3089 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3090 |
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
|
3091 |
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
|
3092 |
|
71990 | 3093 |
lemma bit_word_reverse_iff: |
3094 |
\<open>bit (word_reverse w) n \<longleftrightarrow> n < LENGTH('a) \<and> bit w (LENGTH('a) - Suc n)\<close> |
|
3095 |
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
|
3096 |
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
|
3097 |
(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
|
3098 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3099 |
lemma word_reverse_eq_of_bl_rev_to_bl: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3100 |
\<open>word_reverse w = of_bl (rev (to_bl w))\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3101 |
by (rule bit_word_eqI) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3102 |
(auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3103 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3104 |
lemmas word_reverse_no_def [simp] = |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3105 |
word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3106 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3107 |
lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3108 |
by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3109 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3110 |
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
|
3111 |
by (rule bit_word_eqI) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3112 |
(auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3113 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3114 |
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
|
3115 |
by (metis word_rev_rev) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3116 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3117 |
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
|
3118 |
by simp |
37660 | 3119 |
|
45604 | 3120 |
lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3121 |
|
65268 | 3122 |
lemma nth_sint: |
37660 | 3123 |
fixes w :: "'a::len word" |
70185 | 3124 |
defines "l \<equiv> LENGTH('a)" |
37660 | 3125 |
shows "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" |
3126 |
unfolding sint_uint l_def |
|
65328 | 3127 |
by (auto simp: nth_sbintr word_test_bit_def [symmetric]) |
37660 | 3128 |
|
65328 | 3129 |
lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" |
72000 | 3130 |
apply (simp add: setBit_def bin_sc_eq set_bit_def) |
3131 |
apply transfer |
|
3132 |
apply simp |
|
3133 |
done |
|
3134 |
||
45805 | 3135 |
lemma clearBit_no [simp]: |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54743
diff
changeset
|
3136 |
"clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" |
72000 | 3137 |
apply (simp add: clearBit_def bin_sc_eq unset_bit_def) |
3138 |
apply transfer |
|
3139 |
apply simp |
|
3140 |
done |
|
37660 | 3141 |
|
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
|
3142 |
lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" |
37660 | 3143 |
apply (rule word_bl.Abs_inverse') |
3144 |
apply simp |
|
3145 |
apply (rule word_eqI) |
|
45805 | 3146 |
apply (clarsimp simp add: word_size) |
37660 | 3147 |
apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) |
3148 |
done |
|
3149 |
||
70185 | 3150 |
lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \<longleftrightarrow> m = n \<and> m < LENGTH('a)" |
65328 | 3151 |
by (auto simp: word_test_bit_def word_ubin.eq_norm nth_bintr nth_2p_bin) |
3152 |
||
70185 | 3153 |
lemma nth_w2p: "((2::'a::len word) ^ n) !! m \<longleftrightarrow> m = n \<and> m < LENGTH('a::len)" |
65328 | 3154 |
by (simp add: test_bit_2p [symmetric] word_of_int [symmetric]) |
3155 |
||
3156 |
lemma uint_2p: "(0::'a::len word) < 2 ^ n \<Longrightarrow> uint (2 ^ n::'a::len word) = 2 ^ n" |
|
37660 | 3157 |
apply (unfold word_arith_power_alt) |
70185 | 3158 |
apply (case_tac "LENGTH('a)") |
37660 | 3159 |
apply clarsimp |
3160 |
apply (case_tac "nat") |
|
3161 |
apply clarsimp |
|
3162 |
apply (case_tac "n") |
|
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3163 |
apply clarsimp |
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3164 |
apply clarsimp |
37660 | 3165 |
apply (drule word_gt_0 [THEN iffD1]) |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
3166 |
apply (safe intro!: word_eqI) |
65328 | 3167 |
apply (auto simp add: nth_2p_bin) |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
3168 |
apply (erule notE) |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
3169 |
apply (simp (no_asm_use) add: uint_word_of_int word_size) |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54225
diff
changeset
|
3170 |
apply (subst mod_pos_pos_trivial) |
65328 | 3171 |
apply simp |
3172 |
apply (rule power_strict_increasing) |
|
3173 |
apply simp_all |
|
37660 | 3174 |
done |
3175 |
||
65268 | 3176 |
lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
3177 |
by (induct n) (simp_all add: wi_hom_syms) |
37660 | 3178 |
|
65328 | 3179 |
lemma bang_is_le: "x !! m \<Longrightarrow> 2 ^ m \<le> x" |
3180 |
for x :: "'a::len word" |
|
71997 | 3181 |
apply (rule xtrans(3)) |
65328 | 3182 |
apply (rule_tac [2] y = "x" in le_word_or2) |
37660 | 3183 |
apply (rule word_eqI) |
3184 |
apply (auto simp add: word_ao_nth nth_w2p word_size) |
|
3185 |
done |
|
3186 |
||
70190 | 3187 |
lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\<or>) (rev (to_bl x)) (rev (to_bl y))" |
70193 | 3188 |
by (simp add: zip_rev bl_word_or rev_map) |
70190 | 3189 |
|
3190 |
lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\<and>) (rev (to_bl x)) (rev (to_bl y))" |
|
70193 | 3191 |
by (simp add: zip_rev bl_word_and rev_map) |
70190 | 3192 |
|
3193 |
lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\<noteq>) (rev (to_bl x)) (rev (to_bl y))" |
|
70193 | 3194 |
by (simp add: zip_rev bl_word_xor rev_map) |
70190 | 3195 |
|
3196 |
lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" |
|
3197 |
by (simp add: bl_word_not rev_map) |
|
3198 |
||
37660 | 3199 |
|
70192 | 3200 |
subsection \<open>Bit comprehension\<close> |
3201 |
||
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
|
3202 |
instantiation word :: (len) bit_comprehension |
70192 | 3203 |
begin |
3204 |
||
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3205 |
definition word_set_bits_def: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3206 |
\<open>(BITS n. P n) = (horner_sum of_bool 2 (map P [0..<LENGTH('a)]) :: 'a word)\<close> |
70192 | 3207 |
|
3208 |
instance .. |
|
3209 |
||
3210 |
end |
|
3211 |
||
71990 | 3212 |
lemma bit_set_bits_word_iff: |
3213 |
\<open>bit (set_bits P :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> P n\<close> |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3214 |
by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3215 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3216 |
lemma set_bits_bit_eq: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3217 |
\<open>set_bits (bit w) = w\<close> 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
|
3218 |
by (rule bit_word_eqI) (auto simp add: bit_set_bits_word_iff bit_imp_le_length) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3219 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3220 |
lemma set_bits_K_False [simp]: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3221 |
\<open>set_bits (\<lambda>_. False) = (0 :: 'a :: len word)\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3222 |
by (rule bit_word_eqI) (simp add: bit_set_bits_word_iff) |
71990 | 3223 |
|
70192 | 3224 |
lemmas of_nth_def = word_set_bits_def (* FIXME duplicate *) |
3225 |
||
3226 |
interpretation test_bit: |
|
3227 |
td_ext |
|
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
|
3228 |
"(!!) :: 'a::len word \<Rightarrow> nat \<Rightarrow> bool" |
70192 | 3229 |
set_bits |
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
|
3230 |
"{f. \<forall>i. f i \<longrightarrow> i < LENGTH('a::len)}" |
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
|
3231 |
"(\<lambda>h i. h i \<and> i < LENGTH('a::len))" |
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3232 |
by standard |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3233 |
(auto simp add: test_bit_word_eq bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) |
70192 | 3234 |
|
3235 |
lemmas td_nth = test_bit.td_thm |
|
3236 |
||
3237 |
||
61799 | 3238 |
subsection \<open>Shifting, Rotating, and Splitting Words\<close> |
37660 | 3239 |
|
71986 | 3240 |
lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (2 * w)" |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3241 |
unfolding shiftl1_def |
71986 | 3242 |
apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm) |
71947 | 3243 |
apply (simp add: mod_mult_right_eq take_bit_eq_mod) |
37660 | 3244 |
done |
3245 |
||
65328 | 3246 |
lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3247 |
unfolding word_numeral_alt shiftl1_wi by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3248 |
|
65328 | 3249 |
lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3250 |
unfolding word_neg_numeral_alt shiftl1_wi by simp |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3251 |
|
37660 | 3252 |
lemma shiftl1_0 [simp] : "shiftl1 0 = 0" |
65328 | 3253 |
by (simp add: shiftl1_def) |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3254 |
|
71986 | 3255 |
lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" |
71997 | 3256 |
by (fact shiftl1_def) |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3257 |
|
71986 | 3258 |
lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" |
3259 |
by (simp add: shiftl1_def wi_hom_syms) |
|
37660 | 3260 |
|
45995
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents:
45958
diff
changeset
|
3261 |
lemma shiftr1_0 [simp]: "shiftr1 0 = 0" |
65328 | 3262 |
by (simp add: shiftr1_def) |
45995
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents:
45958
diff
changeset
|
3263 |
|
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents:
45958
diff
changeset
|
3264 |
lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" |
65328 | 3265 |
by (simp add: sshiftr1_def) |
3266 |
||
3267 |
lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" |
|
3268 |
by (simp add: sshiftr1_def) |
|
3269 |
||
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
|
3270 |
lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" |
65328 | 3271 |
by (induct n) (auto simp: shiftl_def) |
3272 |
||
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
|
3273 |
lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" |
65328 | 3274 |
by (induct n) (auto simp: shiftr_def) |
3275 |
||
3276 |
lemma sshiftr_0 [simp]: "0 >>> n = 0" |
|
3277 |
by (induct n) (auto simp: sshiftr_def) |
|
3278 |
||
3279 |
lemma sshiftr_n1 [simp]: "-1 >>> n = -1" |
|
3280 |
by (induct n) (auto simp: sshiftr_def) |
|
3281 |
||
3282 |
lemma nth_shiftl1: "shiftl1 w !! n \<longleftrightarrow> n < size w \<and> n > 0 \<and> w !! (n - 1)" |
|
37660 | 3283 |
apply (unfold shiftl1_def word_test_bit_def) |
3284 |
apply (simp add: nth_bintr word_ubin.eq_norm word_size) |
|
3285 |
apply (cases n) |
|
71986 | 3286 |
apply (simp_all add: bit_Suc) |
37660 | 3287 |
done |
3288 |
||
65328 | 3289 |
lemma nth_shiftl': "(w << m) !! n \<longleftrightarrow> n < size w \<and> n >= m \<and> w !! (n - m)" |
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
|
3290 |
for w :: "'a::len word" |
37660 | 3291 |
apply (unfold shiftl_def) |
65328 | 3292 |
apply (induct m arbitrary: n) |
37660 | 3293 |
apply (force elim!: test_bit_size) |
3294 |
apply (clarsimp simp add : nth_shiftl1 word_size) |
|
3295 |
apply arith |
|
3296 |
done |
|
3297 |
||
65268 | 3298 |
lemmas nth_shiftl = nth_shiftl' [unfolded word_size] |
37660 | 3299 |
|
3300 |
lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" |
|
71949 | 3301 |
apply (auto simp add: shiftr1_def word_test_bit_def word_ubin.eq_norm bit_take_bit_iff bit_Suc) |
3302 |
apply (metis (no_types, hide_lams) add_Suc_right add_diff_cancel_left' bit_Suc diff_is_0_eq' le_Suc_ex less_imp_le linorder_not_le test_bit_bin word_test_bit_def) |
|
37660 | 3303 |
done |
3304 |
||
65328 | 3305 |
lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" |
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
|
3306 |
for w :: "'a::len word" |
37660 | 3307 |
apply (unfold shiftr_def) |
65328 | 3308 |
apply (induct "m" arbitrary: n) |
3309 |
apply (auto simp add: nth_shiftr1) |
|
37660 | 3310 |
done |
65268 | 3311 |
|
67408 | 3312 |
text \<open> |
3313 |
see paper page 10, (1), (2), \<open>shiftr1_def\<close> is of the form of (1), |
|
3314 |
where \<open>f\<close> (ie \<open>bin_rest\<close>) takes normal arguments to normal results, |
|
3315 |
thus we get (2) from (1) |
|
3316 |
\<close> |
|
37660 | 3317 |
|
65268 | 3318 |
lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" |
37660 | 3319 |
apply (unfold shiftr1_def word_ubin.eq_norm bin_rest_trunc_i) |
3320 |
apply (subst bintr_uint [symmetric, OF order_refl]) |
|
3321 |
apply (simp only : bintrunc_bintrunc_l) |
|
65268 | 3322 |
apply simp |
37660 | 3323 |
done |
3324 |
||
71990 | 3325 |
lemma bit_sshiftr1_iff: |
3326 |
\<open>bit (sshiftr1 w) n \<longleftrightarrow> bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\<close> |
|
3327 |
for w :: \<open>'a::len word\<close> |
|
3328 |
apply (cases \<open>LENGTH('a)\<close>) |
|
3329 |
apply simp |
|
3330 |
apply (simp add: sshiftr1_def bit_word_of_int_iff bit_sint_iff flip: bit_Suc) |
|
3331 |
apply transfer apply auto |
|
3332 |
done |
|
3333 |
||
3334 |
lemma bit_sshiftr_word_iff: |
|
3335 |
\<open>bit (w >>> m) n \<longleftrightarrow> bit w (if LENGTH('a) - m \<le> n \<and> n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\<close> |
|
3336 |
for w :: \<open>'a::len word\<close> |
|
3337 |
apply (cases \<open>LENGTH('a)\<close>) |
|
3338 |
apply simp |
|
3339 |
apply (simp only:) |
|
3340 |
apply (induction m arbitrary: n) |
|
3341 |
apply (auto simp add: sshiftr_def bit_sshiftr1_iff not_le less_diff_conv) |
|
3342 |
done |
|
3343 |
||
65328 | 3344 |
lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" |
37660 | 3345 |
apply (unfold sshiftr1_def word_test_bit_def) |
71949 | 3346 |
apply (simp add: nth_bintr word_ubin.eq_norm bit_Suc [symmetric] word_size) |
3347 |
apply (simp add: nth_bintr uint_sint) |
|
37660 | 3348 |
apply (auto simp add: bin_nth_sint) |
3349 |
done |
|
3350 |
||
65268 | 3351 |
lemma nth_sshiftr [rule_format] : |
65328 | 3352 |
"\<forall>n. sshiftr w m !! n = |
3353 |
(n < size w \<and> (if n + m \<ge> size w then w !! (size w - 1) else w !! (n + m)))" |
|
37660 | 3354 |
apply (unfold sshiftr_def) |
65328 | 3355 |
apply (induct_tac m) |
37660 | 3356 |
apply (simp add: test_bit_bl) |
3357 |
apply (clarsimp simp add: nth_sshiftr1 word_size) |
|
3358 |
apply safe |
|
3359 |
apply arith |
|
3360 |
apply arith |
|
3361 |
apply (erule thin_rl) |
|
3362 |
apply (case_tac n) |
|
3363 |
apply safe |
|
3364 |
apply simp |
|
3365 |
apply simp |
|
3366 |
apply (erule thin_rl) |
|
3367 |
apply (case_tac n) |
|
3368 |
apply safe |
|
3369 |
apply simp |
|
3370 |
apply simp |
|
3371 |
apply arith+ |
|
3372 |
done |
|
65268 | 3373 |
|
37660 | 3374 |
lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" |
71945
4b1264316270
replaced operation with weak abstraction by input abbreviation
haftmann
parents:
71944
diff
changeset
|
3375 |
apply (unfold shiftr1_def) |
37660 | 3376 |
apply (rule word_uint.Abs_inverse) |
3377 |
apply (simp add: uints_num pos_imp_zdiv_nonneg_iff) |
|
71997 | 3378 |
apply (metis uint_lt2p uint_shiftr1) |
37660 | 3379 |
done |
3380 |
||
3381 |
lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" |
|
71945
4b1264316270
replaced operation with weak abstraction by input abbreviation
haftmann
parents:
71944
diff
changeset
|
3382 |
apply (unfold sshiftr1_def) |
37660 | 3383 |
apply (simp add: word_sbin.eq_norm) |
3384 |
apply (rule trans) |
|
3385 |
defer |
|
3386 |
apply (subst word_sbin.norm_Rep [symmetric]) |
|
3387 |
apply (rule refl) |
|
3388 |
apply (subst word_sbin.norm_Rep [symmetric]) |
|
3389 |
apply (unfold One_nat_def) |
|
3390 |
apply (rule sbintrunc_rest) |
|
3391 |
done |
|
3392 |
||
3393 |
lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" |
|
3394 |
apply (unfold shiftr_def) |
|
65328 | 3395 |
apply (induct n) |
37660 | 3396 |
apply simp |
65328 | 3397 |
apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) |
37660 | 3398 |
done |
3399 |
||
3400 |
lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" |
|
3401 |
apply (unfold sshiftr_def) |
|
65328 | 3402 |
apply (induct n) |
37660 | 3403 |
apply simp |
65328 | 3404 |
apply (simp add: sshiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) |
37660 | 3405 |
done |
3406 |
||
71990 | 3407 |
lemma bit_bshiftr1_iff: |
3408 |
\<open>bit (bshiftr1 b w) n \<longleftrightarrow> b \<and> n = LENGTH('a) - 1 \<or> bit w (Suc n)\<close> |
|
3409 |
for w :: \<open>'a::len word\<close> |
|
3410 |
apply (cases \<open>LENGTH('a)\<close>) |
|
3411 |
apply simp |
|
3412 |
apply (simp add: bshiftr1_def bit_of_bl_iff nth_append not_less rev_nth nth_butlast nth_to_bl) |
|
3413 |
apply (use bit_imp_le_length in fastforce) |
|
3414 |
done |
|
3415 |
||
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
3416 |
|
61799 | 3417 |
subsubsection \<open>shift functions in terms of lists of bools\<close> |
37660 | 3418 |
|
71997 | 3419 |
lemma bshiftr1_numeral [simp]: |
3420 |
\<open>bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\<close> |
|
3421 |
by (simp add: bshiftr1_def) |
|
37660 | 3422 |
|
3423 |
lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" |
|
3424 |
unfolding bshiftr1_def by (rule word_bl.Abs_inverse) simp |
|
3425 |
||
3426 |
lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" |
|
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3427 |
by (simp add: of_bl_def bl_to_bin_append) |
37660 | 3428 |
|
65363 | 3429 |
lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" |
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
|
3430 |
for w :: "'a::len word" |
37660 | 3431 |
proof - |
65328 | 3432 |
have "shiftl1 w = shiftl1 (of_bl (to_bl w))" |
3433 |
by simp |
|
3434 |
also have "\<dots> = of_bl (to_bl w @ [False])" |
|
3435 |
by (rule shiftl1_of_bl) |
|
37660 | 3436 |
finally show ?thesis . |
3437 |
qed |
|
3438 |
||
65328 | 3439 |
lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" |
3440 |
for w :: "'a::len word" |
|
3441 |
by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) |
|
37660 | 3442 |
|
67408 | 3443 |
\<comment> \<open>Generalized version of \<open>bl_shiftl1\<close>. Maybe this one should replace it?\<close> |
65328 | 3444 |
lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" |
3445 |
by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) |
|
45807 | 3446 |
|
37660 | 3447 |
lemma shiftr1_bl: "shiftr1 w = of_bl (butlast (to_bl w))" |
3448 |
apply (unfold shiftr1_def uint_bl of_bl_def) |
|
3449 |
apply (simp add: butlast_rest_bin word_size) |
|
3450 |
apply (simp add: bin_rest_trunc [symmetric, unfolded One_nat_def]) |
|
3451 |
done |
|
3452 |
||
65328 | 3453 |
lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" |
3454 |
for w :: "'a::len word" |
|
3455 |
by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) |
|
37660 | 3456 |
|
67408 | 3457 |
\<comment> \<open>Generalized version of \<open>bl_shiftr1\<close>. Maybe this one should replace it?\<close> |
65328 | 3458 |
lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" |
45807 | 3459 |
apply (rule word_bl.Abs_inverse') |
65328 | 3460 |
apply (simp del: butlast.simps) |
45807 | 3461 |
apply (simp add: shiftr1_bl of_bl_def) |
3462 |
done |
|
3463 |
||
65328 | 3464 |
lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" |
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3465 |
apply (rule bit_word_eqI) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3466 |
apply (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) |
37660 | 3467 |
done |
3468 |
||
65328 | 3469 |
lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" |
3470 |
by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) |
|
37660 | 3471 |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3472 |
lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3473 |
by (simp add: shiftl_rev) |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3474 |
|
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3475 |
lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3476 |
by (simp add: rev_shiftl) |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3477 |
|
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3478 |
lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
3479 |
by (simp add: shiftr_rev) |
37660 | 3480 |
|
65328 | 3481 |
lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" |
3482 |
for w :: "'a::len word" |
|
37660 | 3483 |
apply (unfold sshiftr1_def uint_bl word_size) |
3484 |
apply (simp add: butlast_rest_bin word_ubin.eq_norm) |
|
3485 |
apply (simp add: sint_uint) |
|
3486 |
apply (rule nth_equalityI) |
|
3487 |
apply clarsimp |
|
3488 |
apply clarsimp |
|
3489 |
apply (case_tac i) |
|
3490 |
apply (simp_all add: hd_conv_nth length_0_conv [symmetric] |
|
71949 | 3491 |
nth_bin_to_bl bit_Suc [symmetric] nth_sbintr) |
37660 | 3492 |
apply force |
3493 |
apply (rule impI) |
|
3494 |
apply (rule_tac f = "bin_nth (uint w)" in arg_cong) |
|
3495 |
apply simp |
|
3496 |
done |
|
3497 |
||
65328 | 3498 |
lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" |
3499 |
for w :: "'a::len word" |
|
37660 | 3500 |
apply (unfold shiftr_def) |
3501 |
apply (induct n) |
|
3502 |
prefer 2 |
|
3503 |
apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) |
|
3504 |
apply (rule butlast_take [THEN trans]) |
|
65328 | 3505 |
apply (auto simp: word_size) |
37660 | 3506 |
done |
3507 |
||
65328 | 3508 |
lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" |
3509 |
for w :: "'a::len word" |
|
37660 | 3510 |
apply (unfold sshiftr_def) |
3511 |
apply (induct n) |
|
3512 |
prefer 2 |
|
3513 |
apply (simp add: drop_Suc bl_sshiftr1 butlast_drop [symmetric]) |
|
3514 |
apply (rule butlast_take [THEN trans]) |
|
65328 | 3515 |
apply (auto simp: word_size) |
37660 | 3516 |
done |
3517 |
||
65328 | 3518 |
lemma take_shiftr: "n \<le> size w \<Longrightarrow> take n (to_bl (w >> n)) = replicate n False" |
37660 | 3519 |
apply (unfold shiftr_def) |
3520 |
apply (induct n) |
|
3521 |
prefer 2 |
|
45807 | 3522 |
apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) |
37660 | 3523 |
apply (rule take_butlast [THEN trans]) |
65328 | 3524 |
apply (auto simp: word_size) |
37660 | 3525 |
done |
3526 |
||
3527 |
lemma take_sshiftr' [rule_format] : |
|
65328 | 3528 |
"n \<le> size w \<longrightarrow> hd (to_bl (w >>> n)) = hd (to_bl w) \<and> |
65268 | 3529 |
take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" |
65328 | 3530 |
for w :: "'a::len word" |
37660 | 3531 |
apply (unfold sshiftr_def) |
3532 |
apply (induct n) |
|
3533 |
prefer 2 |
|
3534 |
apply (simp add: bl_sshiftr1) |
|
3535 |
apply (rule impI) |
|
3536 |
apply (rule take_butlast [THEN trans]) |
|
65328 | 3537 |
apply (auto simp: word_size) |
37660 | 3538 |
done |
3539 |
||
45604 | 3540 |
lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] |
3541 |
lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] |
|
37660 | 3542 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3543 |
lemma atd_lem: "take n xs = t \<Longrightarrow> drop n xs = d \<Longrightarrow> xs = t @ d" |
37660 | 3544 |
by (auto intro: append_take_drop_id [symmetric]) |
3545 |
||
3546 |
lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] |
|
3547 |
lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] |
|
3548 |
||
3549 |
lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" |
|
65328 | 3550 |
by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) |
3551 |
||
3552 |
lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" |
|
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
|
3553 |
for w :: "'a::len word" |
37660 | 3554 |
proof - |
65328 | 3555 |
have "w << n = of_bl (to_bl w) << n" |
3556 |
by simp |
|
3557 |
also have "\<dots> = of_bl (to_bl w @ replicate n False)" |
|
3558 |
by (rule shiftl_of_bl) |
|
37660 | 3559 |
finally show ?thesis . |
3560 |
qed |
|
3561 |
||
71997 | 3562 |
lemma shiftl_numeral [simp]: |
3563 |
\<open>numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\<close> |
|
3564 |
by (fact shiftl_word_eq) |
|
37660 | 3565 |
|
65328 | 3566 |
lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" |
37660 | 3567 |
by (simp add: shiftl_bl word_rep_drop word_size) |
3568 |
||
65328 | 3569 |
lemma shiftl_zero_size: "size x \<le> n \<Longrightarrow> x << n = 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
|
3570 |
for x :: "'a::len word" |
37660 | 3571 |
apply (unfold word_size) |
3572 |
apply (rule word_eqI) |
|
3573 |
apply (clarsimp simp add: shiftl_bl word_size test_bit_of_bl nth_append) |
|
3574 |
done |
|
3575 |
||
67408 | 3576 |
\<comment> \<open>note -- the following results use \<open>'a::len word < number_ring\<close>\<close> |
65268 | 3577 |
|
65328 | 3578 |
lemma shiftl1_2t: "shiftl1 w = 2 * w" |
3579 |
for w :: "'a::len word" |
|
71986 | 3580 |
by (simp add: shiftl1_def wi_hom_mult [symmetric]) |
37660 | 3581 |
|
65328 | 3582 |
lemma shiftl1_p: "shiftl1 w = w + w" |
3583 |
for w :: "'a::len word" |
|
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3584 |
by (simp add: shiftl1_2t) |
37660 | 3585 |
|
65328 | 3586 |
lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" |
3587 |
for w :: "'a::len word" |
|
3588 |
by (induct n) (auto simp: shiftl_def shiftl1_2t) |
|
37660 | 3589 |
|
3590 |
lemma shiftr1_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
|
3591 |
"(shiftr1 (numeral w) :: 'a::len word) = |
70185 | 3592 |
word_of_int (bin_rest (bintrunc (LENGTH('a)) (numeral w)))" |
65328 | 3593 |
unfolding shiftr1_def word_numeral_alt by (simp add: word_ubin.eq_norm) |
46962
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
3594 |
|
5bdcdb28be83
make more word theorems respect int/bin distinction
huffman
parents:
46656
diff
changeset
|
3595 |
lemma sshiftr1_sbintr [simp]: |
65268 | 3596 |
"(sshiftr1 (numeral w) :: 'a::len word) = |
70185 | 3597 |
word_of_int (bin_rest (sbintrunc (LENGTH('a) - 1) (numeral w)))" |
65328 | 3598 |
unfolding sshiftr1_def word_numeral_alt by (simp add: word_sbin.eq_norm) |
37660 | 3599 |
|
71997 | 3600 |
text \<open>TODO: rules for \<^term>\<open>- (numeral n)\<close>\<close> |
3601 |
||
3602 |
lemma drop_bit_word_numeral [simp]: |
|
3603 |
\<open>drop_bit (numeral n) (numeral k) = |
|
3604 |
(word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close> |
|
3605 |
by transfer simp |
|
3606 |
||
3607 |
lemma shiftr_numeral [simp]: |
|
3608 |
\<open>(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\<close> |
|
3609 |
by (fact shiftr_word_eq) |
|
3610 |
||
3611 |
lemma sshiftr_numeral [simp]: |
|
3612 |
\<open>(numeral k >>> numeral n :: 'a::len word) = |
|
3613 |
word_of_int (drop_bit (numeral n) (sbintrunc (LENGTH('a) - 1) (numeral k)))\<close> |
|
37660 | 3614 |
apply (rule word_eqI) |
71997 | 3615 |
apply (cases \<open>LENGTH('a)\<close>) |
3616 |
apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr nth_sbintr not_le not_less less_Suc_eq_le ac_simps) |
|
37660 | 3617 |
done |
3618 |
||
45811 | 3619 |
lemma shiftr1_bl_of: |
70185 | 3620 |
"length bl \<le> LENGTH('a) \<Longrightarrow> |
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
|
3621 |
shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" |
65328 | 3622 |
by (clarsimp simp: shiftr1_def of_bl_def butlast_rest_bl2bin word_ubin.eq_norm trunc_bl2bin) |
37660 | 3623 |
|
45811 | 3624 |
lemma shiftr_bl_of: |
70185 | 3625 |
"length bl \<le> LENGTH('a) \<Longrightarrow> |
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
|
3626 |
(of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" |
37660 | 3627 |
apply (unfold shiftr_def) |
3628 |
apply (induct n) |
|
3629 |
apply clarsimp |
|
3630 |
apply clarsimp |
|
3631 |
apply (subst shiftr1_bl_of) |
|
3632 |
apply simp |
|
3633 |
apply (simp add: butlast_take) |
|
3634 |
done |
|
3635 |
||
70185 | 3636 |
lemma shiftr_bl: "x >> n \<equiv> of_bl (take (LENGTH('a) - n) (to_bl 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
|
3637 |
for x :: "'a::len word" |
45811 | 3638 |
using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp |
3639 |
||
65328 | 3640 |
lemma zip_replicate: "n \<ge> length ys \<Longrightarrow> zip (replicate n x) ys = map (\<lambda>y. (x, y)) ys" |
3641 |
apply (induct ys arbitrary: n) |
|
3642 |
apply simp_all |
|
3643 |
apply (case_tac n) |
|
3644 |
apply simp_all |
|
57492
74bf65a1910a
Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
56979
diff
changeset
|
3645 |
done |
74bf65a1910a
Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
56979
diff
changeset
|
3646 |
|
37660 | 3647 |
lemma align_lem_or [rule_format] : |
65328 | 3648 |
"\<forall>x m. length x = n + m \<longrightarrow> length y = n + m \<longrightarrow> |
3649 |
drop m x = replicate n False \<longrightarrow> take m y = replicate m False \<longrightarrow> |
|
67399 | 3650 |
map2 (|) x y = take m x @ drop m y" |
65328 | 3651 |
apply (induct y) |
37660 | 3652 |
apply force |
3653 |
apply clarsimp |
|
65328 | 3654 |
apply (case_tac x) |
3655 |
apply force |
|
3656 |
apply (case_tac m) |
|
3657 |
apply auto |
|
59807 | 3658 |
apply (drule_tac t="length xs" for xs in sym) |
70193 | 3659 |
apply (auto simp: zip_replicate o_def) |
37660 | 3660 |
done |
3661 |
||
3662 |
lemma align_lem_and [rule_format] : |
|
65328 | 3663 |
"\<forall>x m. length x = n + m \<longrightarrow> length y = n + m \<longrightarrow> |
3664 |
drop m x = replicate n False \<longrightarrow> take m y = replicate m False \<longrightarrow> |
|
67399 | 3665 |
map2 (\<and>) x y = replicate (n + m) False" |
65328 | 3666 |
apply (induct y) |
37660 | 3667 |
apply force |
3668 |
apply clarsimp |
|
65328 | 3669 |
apply (case_tac x) |
3670 |
apply force |
|
3671 |
apply (case_tac m) |
|
3672 |
apply auto |
|
59807 | 3673 |
apply (drule_tac t="length xs" for xs in sym) |
70193 | 3674 |
apply (auto simp: zip_replicate o_def map_replicate_const) |
37660 | 3675 |
done |
3676 |
||
45811 | 3677 |
lemma aligned_bl_add_size [OF refl]: |
65328 | 3678 |
"size x - n = m \<Longrightarrow> n \<le> size x \<Longrightarrow> drop m (to_bl x) = replicate n False \<Longrightarrow> |
65268 | 3679 |
take m (to_bl y) = replicate m False \<Longrightarrow> |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3680 |
to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \<open>'a::len word\<close> |
37660 | 3681 |
apply (subgoal_tac "x AND y = 0") |
3682 |
prefer 2 |
|
3683 |
apply (rule word_bl.Rep_eqD) |
|
45805 | 3684 |
apply (simp add: bl_word_and) |
37660 | 3685 |
apply (rule align_lem_and [THEN trans]) |
3686 |
apply (simp_all add: word_size)[5] |
|
3687 |
apply simp |
|
3688 |
apply (subst word_plus_and_or [symmetric]) |
|
3689 |
apply (simp add : bl_word_or) |
|
3690 |
apply (rule align_lem_or) |
|
3691 |
apply (simp_all add: word_size) |
|
3692 |
done |
|
3693 |
||
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
3694 |
|
61799 | 3695 |
subsubsection \<open>Mask\<close> |
37660 | 3696 |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3697 |
lemma minus_1_eq_mask: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3698 |
\<open>- 1 = (Bit_Operations.mask LENGTH('a) :: 'a::len word)\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3699 |
by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff exp_eq_zero_iff) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3700 |
|
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3701 |
lemma mask_eq_mask: |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3702 |
\<open>mask = Bit_Operations.mask\<close> |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3703 |
by (simp add: fun_eq_iff mask_eq_exp_minus_1 mask_def shiftl_word_eq push_bit_eq_mult) |
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3704 |
|
71953 | 3705 |
lemma mask_eq: |
3706 |
\<open>mask n = 2 ^ n - 1\<close> |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3707 |
by (simp add: mask_eq_mask mask_eq_exp_minus_1) |
71953 | 3708 |
|
3709 |
lemma mask_Suc_rec: |
|
3710 |
\<open>mask (Suc n) = 2 * mask n + 1\<close> |
|
3711 |
by (simp add: mask_eq) |
|
3712 |
||
3713 |
context |
|
3714 |
begin |
|
3715 |
||
71990 | 3716 |
qualified lemma bit_mask_iff: |
3717 |
\<open>bit (mask m :: 'a::len word) n \<longleftrightarrow> n < min LENGTH('a) m\<close> |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
3718 |
by (simp add: mask_eq_mask bit_mask_iff exp_eq_zero_iff not_le) |
71953 | 3719 |
|
3720 |
end |
|
3721 |
||
3722 |
lemma nth_mask [simp]: |
|
3723 |
\<open>(mask n :: 'a::len word) !! i \<longleftrightarrow> i < n \<and> i < size (mask n :: 'a word)\<close> |
|
71990 | 3724 |
by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) |
37660 | 3725 |
|
3726 |
lemma mask_bl: "mask n = of_bl (replicate n True)" |
|
3727 |
by (auto simp add : test_bit_of_bl word_size intro: word_eqI) |
|
3728 |
||
58410
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents:
58061
diff
changeset
|
3729 |
lemma mask_bin: "mask n = word_of_int (bintrunc n (- 1))" |
37660 | 3730 |
by (auto simp add: nth_bintr word_size intro: word_eqI) |
3731 |
||
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3732 |
lemma and_mask_bintr: "w AND mask n = word_of_int (bintrunc n (uint w))" |
37660 | 3733 |
apply (rule word_eqI) |
3734 |
apply (simp add: nth_bintr word_size word_ops_nth_size) |
|
3735 |
apply (auto simp add: test_bit_bin) |
|
3736 |
done |
|
3737 |
||
45811 | 3738 |
lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (bintrunc n i)" |
46023
fad87bb608fc
restate some lemmas to respect int/bin distinction
huffman
parents:
46022
diff
changeset
|
3739 |
by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) |
fad87bb608fc
restate some lemmas to respect int/bin distinction
huffman
parents:
46022
diff
changeset
|
3740 |
|
65328 | 3741 |
lemma and_mask_wi': |
3742 |
"word_of_int i AND mask n = (word_of_int (bintrunc (min LENGTH('a) n) i) :: 'a::len word)" |
|
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
3743 |
by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) |
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
3744 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3745 |
lemma and_mask_no: "numeral i AND mask n = word_of_int (bintrunc n (numeral i))" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3746 |
unfolding word_numeral_alt by (rule and_mask_wi) |
37660 | 3747 |
|
3748 |
lemma bl_and_mask': |
|
65268 | 3749 |
"to_bl (w AND mask n :: 'a::len word) = |
70185 | 3750 |
replicate (LENGTH('a) - n) False @ |
3751 |
drop (LENGTH('a) - n) (to_bl w)" |
|
37660 | 3752 |
apply (rule nth_equalityI) |
3753 |
apply simp |
|
3754 |
apply (clarsimp simp add: to_bl_nth word_size) |
|
3755 |
apply (simp add: word_size word_ops_nth_size) |
|
71997 | 3756 |
apply (auto simp add: word_size test_bit_bl nth_append rev_nth) |
37660 | 3757 |
done |
3758 |
||
45811 | 3759 |
lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" |
46001
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
huffman
parents:
46000
diff
changeset
|
3760 |
by (simp only: and_mask_bintr bintrunc_mod2p) |
37660 | 3761 |
|
3762 |
lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" |
|
71997 | 3763 |
by (simp add: and_mask_bintr uint.abs_eq min_def not_le lt2p_lem bintr_lt2p) |
37660 | 3764 |
|
65363 | 3765 |
lemma eq_mod_iff: "0 < n \<Longrightarrow> b = b mod n \<longleftrightarrow> 0 \<le> b \<and> b < n" |
3766 |
for b n :: int |
|
71997 | 3767 |
by auto (metis pos_mod_conj)+ |
37660 | 3768 |
|
65363 | 3769 |
lemma mask_eq_iff: "w AND mask n = w \<longleftrightarrow> uint w < 2 ^ n" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3770 |
apply (simp add: and_mask_bintr) |
37660 | 3771 |
apply (simp add: word_ubin.inverse_norm) |
3772 |
apply (simp add: eq_mod_iff bintrunc_mod2p min_def) |
|
3773 |
apply (fast intro!: lt2p_lem) |
|
3774 |
done |
|
3775 |
||
65328 | 3776 |
lemma and_mask_dvd: "2 ^ n dvd uint w \<longleftrightarrow> w AND mask n = 0" |
37660 | 3777 |
apply (simp add: dvd_eq_mod_eq_0 and_mask_mod_2p) |
65328 | 3778 |
apply (simp add: word_uint.norm_eq_iff [symmetric] word_of_int_homs del: word_of_int_0) |
37660 | 3779 |
apply (subst word_uint.norm_Rep [symmetric]) |
3780 |
apply (simp only: bintrunc_bintrunc_min bintrunc_mod2p [symmetric] min_def) |
|
3781 |
apply auto |
|
3782 |
done |
|
3783 |
||
65328 | 3784 |
lemma and_mask_dvd_nat: "2 ^ n dvd unat w \<longleftrightarrow> w AND mask n = 0" |
37660 | 3785 |
apply (unfold unat_def) |
3786 |
apply (rule trans [OF _ and_mask_dvd]) |
|
65268 | 3787 |
apply (unfold dvd_def) |
3788 |
apply auto |
|
65328 | 3789 |
apply (drule uint_ge_0 [THEN nat_int.Abs_inverse' [simplified], symmetric]) |
3790 |
apply simp |
|
3791 |
apply (simp add: nat_mult_distrib nat_power_eq) |
|
37660 | 3792 |
done |
3793 |
||
65328 | 3794 |
lemma word_2p_lem: "n < size w \<Longrightarrow> w < 2 ^ n = (uint w < 2 ^ n)" |
3795 |
for w :: "'a::len word" |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
3796 |
apply (unfold word_size word_less_alt word_numeral_alt) |
71942 | 3797 |
apply (auto simp add: word_of_int_power_hom word_uint.eq_norm |
65328 | 3798 |
simp del: word_of_int_numeral) |
37660 | 3799 |
done |
3800 |
||
65328 | 3801 |
lemma less_mask_eq: "x < 2 ^ n \<Longrightarrow> x AND mask n = x" |
3802 |
for x :: "'a::len word" |
|
71997 | 3803 |
apply (simp add: and_mask_bintr) |
3804 |
apply transfer |
|
3805 |
apply (simp add: ac_simps) |
|
3806 |
apply (auto simp add: min_def) |
|
3807 |
apply (metis bintrunc_bintrunc_ge mod_pos_pos_trivial mult.commute mult.left_neutral mult_zero_left not_le of_bool_def take_bit_eq_mod take_bit_nonnegative) |
|
37660 | 3808 |
done |
3809 |
||
45604 | 3810 |
lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] |
3811 |
||
3812 |
lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] |
|
37660 | 3813 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3814 |
lemma and_mask_less_size: "n < size x \<Longrightarrow> x AND mask n < 2^n" |
37660 | 3815 |
unfolding word_size by (erule and_mask_less') |
3816 |
||
65328 | 3817 |
lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \<Longrightarrow> c > 0 \<Longrightarrow> x mod c = x AND mask n" |
3818 |
for c x :: "'a::len word" |
|
3819 |
by (auto simp: word_mod_def uint_2p and_mask_mod_2p) |
|
37660 | 3820 |
|
3821 |
lemma mask_eqs: |
|
3822 |
"(a AND mask n) + b AND mask n = a + b AND mask n" |
|
3823 |
"a + (b AND mask n) AND mask n = a + b AND mask n" |
|
3824 |
"(a AND mask n) - b AND mask n = a - b AND mask n" |
|
3825 |
"a - (b AND mask n) AND mask n = a - b AND mask n" |
|
3826 |
"a * (b AND mask n) AND mask n = a * b AND mask n" |
|
3827 |
"(b AND mask n) * a AND mask n = b * a AND mask n" |
|
3828 |
"(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" |
|
3829 |
"(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" |
|
3830 |
"(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" |
|
3831 |
"- (a AND mask n) AND mask n = - a AND mask n" |
|
3832 |
"word_succ (a AND mask n) AND mask n = word_succ a AND mask n" |
|
3833 |
"word_pred (a AND mask n) AND mask n = word_pred a AND mask n" |
|
3834 |
using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] |
|
65328 | 3835 |
by (auto simp: and_mask_wi' word_of_int_homs word.abs_eq_iff bintrunc_mod2p mod_simps) |
3836 |
||
3837 |
lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" |
|
37660 | 3838 |
using word_of_int_Ex [where x=x] |
65328 | 3839 |
by (auto simp: and_mask_wi' word_of_int_power_hom word.abs_eq_iff bintrunc_mod2p mod_simps) |
37660 | 3840 |
|
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
3841 |
lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
3842 |
by (simp add: mask_def word_size shiftl_zero_size) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
3843 |
|
37660 | 3844 |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3845 |
subsubsection \<open>Slices\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3846 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3847 |
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
|
3848 |
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
|
3849 |
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
|
3850 |
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
|
3851 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3852 |
lemma bit_slice1_iff: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3853 |
\<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
|
3854 |
\<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
|
3855 |
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
|
3856 |
by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff exp_eq_zero_iff not_less not_le ac_simps |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3857 |
dest: bit_imp_le_length) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3858 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3859 |
lemma slice1_eq_of_bl: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3860 |
\<open>(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3861 |
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
|
3862 |
proof (rule bit_word_eqI) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3863 |
fix m |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3864 |
assume \<open>m \<le> LENGTH('b)\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3865 |
show \<open>bit (slice1 n w :: 'b::len word) m \<longleftrightarrow> bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3866 |
by (cases \<open>m \<ge> n\<close>; cases \<open>LENGTH('a) \<ge> n\<close>) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3867 |
(auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3868 |
qed |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3869 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3870 |
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
|
3871 |
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
|
3872 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3873 |
lemma bit_slice_iff: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3874 |
\<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
|
3875 |
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
|
3876 |
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
|
3877 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3878 |
lemma slice1_no_bin [simp]: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3879 |
"slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3880 |
by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3881 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3882 |
lemma slice_no_bin [simp]: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3883 |
"slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3884 |
(bin_to_bl (LENGTH('b::len)) (numeral w)))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3885 |
by (simp add: slice_def) (* TODO: neg_numeral *) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3886 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3887 |
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
|
3888 |
unfolding slice1_def by simp |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3889 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3890 |
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
|
3891 |
unfolding slice_def by auto |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3892 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3893 |
lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3894 |
by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3895 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3896 |
lemmas slice_take = slice_take' [unfolded word_size] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3897 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3898 |
\<comment> \<open>shiftr to a word of the same size is just slice, |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3899 |
slice is just shiftr then ucast\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3900 |
lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3901 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3902 |
lemma slice_shiftr: "slice n w = ucast (w >> n)" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3903 |
apply (unfold slice_take shiftr_bl) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3904 |
apply (rule ucast_of_bl_up [symmetric]) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3905 |
apply (simp add: word_size) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3906 |
done |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3907 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3908 |
lemma nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \<and> m < LENGTH('a))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3909 |
by (simp add: slice_shiftr nth_ucast nth_shiftr) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3910 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3911 |
lemma slice1_down_alt': |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3912 |
"sl = slice1 n w \<Longrightarrow> fs = size sl \<Longrightarrow> fs + k = n \<Longrightarrow> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3913 |
to_bl sl = takefill False fs (drop k (to_bl w))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3914 |
by (auto simp: slice1_eq_of_bl word_size of_bl_def uint_bl |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3915 |
word_ubin.eq_norm bl_bin_bl_rep_drop drop_takefill) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3916 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3917 |
lemma slice1_up_alt': |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3918 |
"sl = slice1 n w \<Longrightarrow> fs = size sl \<Longrightarrow> fs = n + k \<Longrightarrow> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3919 |
to_bl sl = takefill False fs (replicate k False @ (to_bl w))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3920 |
apply (unfold slice1_eq_of_bl word_size of_bl_def uint_bl) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3921 |
apply (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop takefill_append [symmetric]) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3922 |
apply (rule_tac f = "\<lambda>k. takefill False (LENGTH('a)) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3923 |
(replicate k False @ bin_to_bl (LENGTH('b)) (uint w))" in arg_cong) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3924 |
apply arith |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3925 |
done |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3926 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3927 |
lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3928 |
lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3929 |
lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3930 |
lemmas slice1_up_alts = |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3931 |
le_add_diff_inverse [symmetric, THEN su1] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3932 |
le_add_diff_inverse2 [symmetric, THEN su1] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3933 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3934 |
lemma ucast_slice1: "ucast w = slice1 (size w) w" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3935 |
by (simp add: slice1_def ucast_bl takefill_same' word_size) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3936 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3937 |
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
|
3938 |
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
|
3939 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3940 |
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
|
3941 |
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
|
3942 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3943 |
lemma slice1_tf_tf': |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3944 |
"to_bl (slice1 n w :: 'a::len word) = |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3945 |
rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3946 |
unfolding slice1_eq_of_bl by (rule word_rev_tf) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3947 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3948 |
lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3949 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3950 |
lemma rev_slice1: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3951 |
"n + k = LENGTH('a) + LENGTH('b) \<Longrightarrow> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3952 |
slice1 n (word_reverse w :: 'b::len word) = |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3953 |
word_reverse (slice1 k w :: 'a::len word)" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3954 |
apply (unfold word_reverse_eq_of_bl_rev_to_bl slice1_tf_tf) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3955 |
apply (rule word_bl.Rep_inverse') |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3956 |
apply (rule rev_swap [THEN iffD1]) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3957 |
apply (rule trans [symmetric]) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3958 |
apply (rule tf_rev) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3959 |
apply (simp add: word_bl.Abs_inverse) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3960 |
apply (simp add: word_bl.Abs_inverse) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3961 |
done |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3962 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3963 |
lemma rev_slice: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3964 |
"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
|
3965 |
slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3966 |
apply (unfold slice_def word_size) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3967 |
apply (rule rev_slice1) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3968 |
apply arith |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3969 |
done |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3970 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3971 |
|
61799 | 3972 |
subsubsection \<open>Revcast\<close> |
37660 | 3973 |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3974 |
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
|
3975 |
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
|
3976 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3977 |
lemma bit_revcast_iff: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3978 |
\<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
|
3979 |
\<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
|
3980 |
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
|
3981 |
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
|
3982 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3983 |
lemma revcast_eq_of_bl: |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3984 |
\<open>(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\<close> |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3985 |
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
|
3986 |
by (simp add: revcast_def slice1_eq_of_bl) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3987 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3988 |
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
|
3989 |
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
|
3990 |
|
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3991 |
lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w |
37660 | 3992 |
|
65268 | 3993 |
lemma to_bl_revcast: |
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
|
3994 |
"to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" |
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3995 |
apply (rule nth_equalityI) |
37660 | 3996 |
apply simp |
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3997 |
apply (cases \<open>LENGTH('a) \<le> LENGTH('b)\<close>) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
3998 |
apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) |
37660 | 3999 |
done |
4000 |
||
65268 | 4001 |
lemma revcast_rev_ucast [OF refl refl refl]: |
4002 |
"cs = [rc, uc] \<Longrightarrow> rc = revcast (word_reverse w) \<Longrightarrow> uc = ucast w \<Longrightarrow> |
|
37660 | 4003 |
rc = word_reverse uc" |
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4004 |
apply auto |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4005 |
apply (rule bit_word_eqI) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4006 |
apply (cases \<open>LENGTH('a) \<le> LENGTH('b)\<close>) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4007 |
apply (simp_all add: bit_revcast_iff bit_word_reverse_iff bit_ucast_iff not_le |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4008 |
bit_imp_le_length) |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4009 |
using bit_imp_le_length apply fastforce |
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4010 |
using bit_imp_le_length apply fastforce |
37660 | 4011 |
done |
4012 |
||
45811 | 4013 |
lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" |
4014 |
using revcast_rev_ucast [of "word_reverse w"] by simp |
|
4015 |
||
4016 |
lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" |
|
4017 |
by (fact revcast_rev_ucast [THEN word_rev_gal']) |
|
4018 |
||
4019 |
lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" |
|
4020 |
by (fact revcast_ucast [THEN word_rev_gal']) |
|
37660 | 4021 |
|
4022 |
||
65328 | 4023 |
text "linking revcast and cast via shift" |
37660 | 4024 |
|
4025 |
lemmas wsst_TYs = source_size target_size word_size |
|
4026 |
||
45811 | 4027 |
lemma revcast_down_uu [OF refl]: |
65328 | 4028 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> rc w = ucast (w >> n)" |
4029 |
for w :: "'a::len word" |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4030 |
apply (simp add: revcast_eq_of_bl) |
37660 | 4031 |
apply (rule word_bl.Rep_inverse') |
4032 |
apply (rule trans, rule ucast_down_drop) |
|
4033 |
prefer 2 |
|
4034 |
apply (rule trans, rule drop_shiftr) |
|
4035 |
apply (auto simp: takefill_alt wsst_TYs) |
|
4036 |
done |
|
4037 |
||
45811 | 4038 |
lemma revcast_down_us [OF refl]: |
65328 | 4039 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> rc w = ucast (w >>> n)" |
4040 |
for w :: "'a::len word" |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4041 |
apply (simp add: revcast_eq_of_bl) |
37660 | 4042 |
apply (rule word_bl.Rep_inverse') |
4043 |
apply (rule trans, rule ucast_down_drop) |
|
4044 |
prefer 2 |
|
4045 |
apply (rule trans, rule drop_sshiftr) |
|
4046 |
apply (auto simp: takefill_alt wsst_TYs) |
|
4047 |
done |
|
4048 |
||
45811 | 4049 |
lemma revcast_down_su [OF refl]: |
65328 | 4050 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> rc w = scast (w >> n)" |
4051 |
for w :: "'a::len word" |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4052 |
apply (simp add: revcast_eq_of_bl) |
37660 | 4053 |
apply (rule word_bl.Rep_inverse') |
4054 |
apply (rule trans, rule scast_down_drop) |
|
4055 |
prefer 2 |
|
4056 |
apply (rule trans, rule drop_shiftr) |
|
4057 |
apply (auto simp: takefill_alt wsst_TYs) |
|
4058 |
done |
|
4059 |
||
45811 | 4060 |
lemma revcast_down_ss [OF refl]: |
65328 | 4061 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> rc w = scast (w >>> n)" |
4062 |
for w :: "'a::len word" |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4063 |
apply (simp add: revcast_eq_of_bl) |
37660 | 4064 |
apply (rule word_bl.Rep_inverse') |
4065 |
apply (rule trans, rule scast_down_drop) |
|
4066 |
prefer 2 |
|
4067 |
apply (rule trans, rule drop_sshiftr) |
|
4068 |
apply (auto simp: takefill_alt wsst_TYs) |
|
4069 |
done |
|
4070 |
||
45811 | 4071 |
(* FIXME: should this also be [OF refl] ? *) |
65268 | 4072 |
lemma cast_down_rev: |
65328 | 4073 |
"uc = ucast \<Longrightarrow> source_size uc = target_size uc + n \<Longrightarrow> uc w = revcast (w << n)" |
4074 |
for w :: "'a::len word" |
|
37660 | 4075 |
apply (unfold shiftl_rev) |
4076 |
apply clarify |
|
4077 |
apply (simp add: revcast_rev_ucast) |
|
4078 |
apply (rule word_rev_gal') |
|
4079 |
apply (rule trans [OF _ revcast_rev_ucast]) |
|
4080 |
apply (rule revcast_down_uu [symmetric]) |
|
4081 |
apply (auto simp add: wsst_TYs) |
|
4082 |
done |
|
4083 |
||
45811 | 4084 |
lemma revcast_up [OF refl]: |
65268 | 4085 |
"rc = revcast \<Longrightarrow> source_size rc + n = target_size rc \<Longrightarrow> |
4086 |
rc w = (ucast w :: 'a::len word) << n" |
|
72027
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
haftmann
parents:
72010
diff
changeset
|
4087 |
apply (simp add: revcast_eq_of_bl) |
37660 | 4088 |
apply (rule word_bl.Rep_inverse') |
4089 |
apply (simp add: takefill_alt) |
|
4090 |
apply (rule bl_shiftl [THEN trans]) |
|
4091 |
apply (subst ucast_up_app) |
|
65328 | 4092 |
apply (auto simp add: wsst_TYs) |
37660 | 4093 |
done |
4094 |
||
65268 | 4095 |
lemmas rc1 = revcast_up [THEN |
37660 | 4096 |
revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] |
65268 | 4097 |
lemmas rc2 = revcast_down_uu [THEN |
37660 | 4098 |
revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] |
4099 |
||
4100 |
lemmas ucast_up = |
|
4101 |
rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] |
|
65268 | 4102 |
lemmas ucast_down = |
37660 | 4103 |
rc2 [simplified rev_shiftr revcast_ucast [symmetric]] |
4104 |
||
65268 | 4105 |
lemmas sym_notr = |
37660 | 4106 |
not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] |
4107 |
||
61799 | 4108 |
\<comment> \<open>problem posed by TPHOLs referee: |
4109 |
criterion for overflow of addition of signed integers\<close> |
|
37660 | 4110 |
|
4111 |
lemma sofl_test: |
|
72000 | 4112 |
\<open>sint x + sint y = sint (x + y) \<longleftrightarrow> |
4113 |
(x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\<close> |
|
4114 |
for x y :: \<open>'a::len word\<close> |
|
4115 |
proof - |
|
4116 |
obtain n where n: \<open>LENGTH('a) = Suc n\<close> |
|
4117 |
by (cases \<open>LENGTH('a)\<close>) simp_all |
|
72010 | 4118 |
have *: \<open>sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \<Longrightarrow> sint x + sint y \<ge> - (2 ^ n)\<close> |
4119 |
\<open>signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \<Longrightarrow> 2 ^ n > sint x + sint y\<close> |
|
4120 |
using signed_take_bit_greater_eq [of \<open>sint x + sint y\<close> n] signed_take_bit_less_eq [of n \<open>sint x + sint y\<close>] |
|
4121 |
by (auto intro: ccontr) |
|
72000 | 4122 |
have \<open>sint x + sint y = sint (x + y) \<longleftrightarrow> |
4123 |
(sint (x + y) < 0 \<longleftrightarrow> sint x < 0) \<or> |
|
4124 |
(sint (x + y) < 0 \<longleftrightarrow> sint y < 0)\<close> |
|
72010 | 4125 |
using sint_range' [of x] sint_range' [of y] |
4126 |
apply (auto simp add: not_less) |
|
4127 |
apply (unfold sint_word_ariths word_sbin.set_iff_norm [symmetric] sints_num) |
|
4128 |
apply (auto simp add: signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) |
|
72000 | 4129 |
done |
4130 |
then show ?thesis |
|
72010 | 4131 |
apply (simp add: word_size shiftr_word_eq drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) |
72000 | 4132 |
apply (simp add: bit_last_iff) |
4133 |
done |
|
4134 |
qed |
|
37660 | 4135 |
|
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
4136 |
lemma shiftr_zero_size: "size x \<le> n \<Longrightarrow> x >> n = 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
|
4137 |
for x :: "'a :: len word" |
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
4138 |
by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
4139 |
|
37660 | 4140 |
|
61799 | 4141 |
subsection \<open>Split and cat\<close> |
37660 | 4142 |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4143 |
lemmas word_split_bin' = word_split_def |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4144 |
lemmas word_cat_bin' = word_cat_def |
37660 | 4145 |
|
4146 |
lemma word_rsplit_no: |
|
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
|
4147 |
"(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = |
70185 | 4148 |
map word_of_int (bin_rsplit (LENGTH('a::len)) |
4149 |
(LENGTH('b), bintrunc (LENGTH('b)) (numeral bin)))" |
|
65336 | 4150 |
by (simp add: word_rsplit_def word_ubin.eq_norm) |
37660 | 4151 |
|
4152 |
lemmas word_rsplit_no_cl [simp] = word_rsplit_no |
|
4153 |
[unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] |
|
4154 |
||
4155 |
lemma test_bit_cat: |
|
65336 | 4156 |
"wc = word_cat a b \<Longrightarrow> wc !! n = (n < size wc \<and> |
37660 | 4157 |
(if n < size b then b !! n else a !! (n - size b)))" |
65336 | 4158 |
apply (auto simp: word_cat_bin' test_bit_bin word_ubin.eq_norm nth_bintr bin_nth_cat word_size) |
37660 | 4159 |
apply (erule bin_nth_uint_imp) |
4160 |
done |
|
4161 |
||
4162 |
lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" |
|
65336 | 4163 |
by (simp add: of_bl_def to_bl_def word_cat_bin' bl_to_bin_app_cat) |
37660 | 4164 |
|
4165 |
lemma of_bl_append: |
|
65268 | 4166 |
"(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" |
65336 | 4167 |
apply (simp add: of_bl_def bl_to_bin_app_cat bin_cat_num) |
46009 | 4168 |
apply (simp add: word_of_int_power_hom [symmetric] word_of_int_hom_syms) |
37660 | 4169 |
done |
4170 |
||
65336 | 4171 |
lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" |
4172 |
by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) |
|
4173 |
||
4174 |
lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" |
|
4175 |
by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) |
|
4176 |
||
4177 |
lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" |
|
45805 | 4178 |
by (cases x) simp_all |
37660 | 4179 |
|
65336 | 4180 |
lemma split_uint_lem: "bin_split n (uint w) = (a, b) \<Longrightarrow> |
70185 | 4181 |
a = bintrunc (LENGTH('a) - n) a \<and> b = bintrunc (LENGTH('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
|
4182 |
for w :: "'a::len word" |
37660 | 4183 |
apply (frule word_ubin.norm_Rep [THEN ssubst]) |
4184 |
apply (drule bin_split_trunc1) |
|
4185 |
apply (drule sym [THEN trans]) |
|
65336 | 4186 |
apply assumption |
37660 | 4187 |
apply safe |
4188 |
done |
|
4189 |
||
65268 | 4190 |
lemma word_split_bl': |
4191 |
"std = size c - size b \<Longrightarrow> (word_split c = (a, b)) \<Longrightarrow> |
|
65336 | 4192 |
(a = of_bl (take std (to_bl c)) \<and> b = of_bl (drop std (to_bl c)))" |
37660 | 4193 |
apply (unfold word_split_bin') |
4194 |
apply safe |
|
4195 |
defer |
|
4196 |
apply (clarsimp split: prod.splits) |
|
71947 | 4197 |
apply (metis of_bl_drop' ucast_bl ucast_def word_size word_size_bl) |
57492
74bf65a1910a
Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents:
56979
diff
changeset
|
4198 |
apply hypsubst_thin |
37660 | 4199 |
apply (drule word_ubin.norm_Rep [THEN ssubst]) |
65336 | 4200 |
apply (simp add: of_bl_def bl2bin_drop word_size |
4201 |
word_ubin.norm_eq_iff [symmetric] min_def del: word_ubin.norm_Rep) |
|
37660 | 4202 |
apply (clarsimp split: prod.splits) |
70185 | 4203 |
apply (cases "LENGTH('a) \<ge> LENGTH('b)") |
71944 | 4204 |
apply (simp_all add: not_le) |
4205 |
defer |
|
4206 |
apply (simp add: drop_bit_eq_div lt2p_lem) |
|
37660 | 4207 |
apply (simp add : of_bl_def to_bl_def) |
71944 | 4208 |
apply (subst bin_to_bl_drop_bit [symmetric]) |
4209 |
apply (subst diff_add) |
|
71947 | 4210 |
apply (simp_all add: take_bit_drop_bit) |
37660 | 4211 |
done |
4212 |
||
65268 | 4213 |
lemma word_split_bl: "std = size c - size b \<Longrightarrow> |
65336 | 4214 |
(a = of_bl (take std (to_bl c)) \<and> b = of_bl (drop std (to_bl c))) \<longleftrightarrow> |
37660 | 4215 |
word_split c = (a, b)" |
4216 |
apply (rule iffI) |
|
4217 |
defer |
|
4218 |
apply (erule (1) word_split_bl') |
|
4219 |
apply (case_tac "word_split c") |
|
65336 | 4220 |
apply (auto simp add: word_size) |
37660 | 4221 |
apply (frule word_split_bl' [rotated]) |
65336 | 4222 |
apply (auto simp add: word_size) |
37660 | 4223 |
done |
4224 |
||
4225 |
lemma word_split_bl_eq: |
|
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
|
4226 |
"(word_split c :: ('c::len word \<times> 'd::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
|
4227 |
(of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), |
70185 | 4228 |
of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" |
65336 | 4229 |
for c :: "'a::len word" |
37660 | 4230 |
apply (rule word_split_bl [THEN iffD1]) |
65336 | 4231 |
apply (unfold word_size) |
4232 |
apply (rule refl conjI)+ |
|
37660 | 4233 |
done |
4234 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
4235 |
\<comment> \<open>keep quantifiers for use in simplification\<close> |
37660 | 4236 |
lemma test_bit_split': |
65336 | 4237 |
"word_split c = (a, b) \<longrightarrow> |
4238 |
(\<forall>n m. |
|
4239 |
b !! n = (n < size b \<and> c !! n) \<and> |
|
4240 |
a !! m = (m < size a \<and> c !! (m + size b)))" |
|
37660 | 4241 |
apply (unfold word_split_bin' test_bit_bin) |
4242 |
apply (clarify) |
|
4243 |
apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits) |
|
71949 | 4244 |
apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq ac_simps bin_nth_uint_imp) |
37660 | 4245 |
done |
4246 |
||
4247 |
lemma test_bit_split: |
|
4248 |
"word_split c = (a, b) \<Longrightarrow> |
|
65336 | 4249 |
(\<forall>n::nat. b !! n \<longleftrightarrow> n < size b \<and> c !! n) \<and> |
4250 |
(\<forall>m::nat. a !! m \<longleftrightarrow> m < size a \<and> c !! (m + size b))" |
|
37660 | 4251 |
by (simp add: test_bit_split') |
4252 |
||
65336 | 4253 |
lemma test_bit_split_eq: |
4254 |
"word_split c = (a, b) \<longleftrightarrow> |
|
4255 |
((\<forall>n::nat. b !! n = (n < size b \<and> c !! n)) \<and> |
|
4256 |
(\<forall>m::nat. a !! m = (m < size a \<and> c !! (m + size b))))" |
|
37660 | 4257 |
apply (rule_tac iffI) |
4258 |
apply (rule_tac conjI) |
|
4259 |
apply (erule test_bit_split [THEN conjunct1]) |
|
4260 |
apply (erule test_bit_split [THEN conjunct2]) |
|
4261 |
apply (case_tac "word_split c") |
|
4262 |
apply (frule test_bit_split) |
|
4263 |
apply (erule trans) |
|
65336 | 4264 |
apply (fastforce intro!: word_eqI simp add: word_size) |
37660 | 4265 |
done |
4266 |
||
65268 | 4267 |
\<comment> \<open>this odd result is analogous to \<open>ucast_id\<close>, |
61799 | 4268 |
result to the length given by the result type\<close> |
37660 | 4269 |
|
4270 |
lemma word_cat_id: "word_cat a b = b" |
|
65336 | 4271 |
by (simp add: word_cat_bin' word_ubin.inverse_norm) |
37660 | 4272 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
4273 |
\<comment> \<open>limited hom result\<close> |
37660 | 4274 |
lemma word_cat_hom: |
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
|
4275 |
"LENGTH('a::len) \<le> LENGTH('b::len) + LENGTH('c::len) \<Longrightarrow> |
65336 | 4276 |
(word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = |
4277 |
word_of_int (bin_cat w (size b) (uint b))" |
|
4278 |
by (auto simp: word_cat_def word_size word_ubin.norm_eq_iff [symmetric] |
|
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54854
diff
changeset
|
4279 |
word_ubin.eq_norm bintr_cat min.absorb1) |
65336 | 4280 |
|
4281 |
lemma word_cat_split_alt: "size w \<le> size u + size v \<Longrightarrow> word_split w = (u, v) \<Longrightarrow> word_cat u v = w" |
|
37660 | 4282 |
apply (rule word_eqI) |
4283 |
apply (drule test_bit_split) |
|
4284 |
apply (clarsimp simp add : test_bit_cat word_size) |
|
4285 |
apply safe |
|
4286 |
apply arith |
|
4287 |
done |
|
4288 |
||
45604 | 4289 |
lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] |
37660 | 4290 |
|
4291 |
||
61799 | 4292 |
subsubsection \<open>Split and slice\<close> |
37660 | 4293 |
|
65336 | 4294 |
lemma split_slices: "word_split w = (u, v) \<Longrightarrow> u = slice (size v) w \<and> v = slice 0 w" |
37660 | 4295 |
apply (drule test_bit_split) |
4296 |
apply (rule conjI) |
|
4297 |
apply (rule word_eqI, clarsimp simp: nth_slice word_size)+ |
|
4298 |
done |
|
4299 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4300 |
lemma slice_cat1 [OF refl]: |
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4301 |
"wc = word_cat a b \<Longrightarrow> size wc >= size a + size b \<Longrightarrow> slice (size b) wc = a" |
37660 | 4302 |
apply safe |
4303 |
apply (rule word_eqI) |
|
4304 |
apply (simp add: nth_slice test_bit_cat word_size) |
|
4305 |
done |
|
4306 |
||
4307 |
lemmas slice_cat2 = trans [OF slice_id word_cat_id] |
|
4308 |
||
4309 |
lemma cat_slices: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4310 |
"a = slice n c \<Longrightarrow> b = slice 0 c \<Longrightarrow> n = size b \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4311 |
size a + size b >= size c \<Longrightarrow> word_cat a b = c" |
37660 | 4312 |
apply safe |
4313 |
apply (rule word_eqI) |
|
4314 |
apply (simp add: nth_slice test_bit_cat word_size) |
|
4315 |
apply safe |
|
4316 |
apply arith |
|
4317 |
done |
|
4318 |
||
4319 |
lemma word_split_cat_alt: |
|
65336 | 4320 |
"w = word_cat u v \<Longrightarrow> size u + size v \<le> size w \<Longrightarrow> word_split w = (u, v)" |
59807 | 4321 |
apply (case_tac "word_split w") |
37660 | 4322 |
apply (rule trans, assumption) |
4323 |
apply (drule test_bit_split) |
|
4324 |
apply safe |
|
4325 |
apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+ |
|
4326 |
done |
|
4327 |
||
4328 |
lemmas word_cat_bl_no_bin [simp] = |
|
65336 | 4329 |
word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4330 |
for a b (* FIXME: negative numerals, 0 and 1 *) |
37660 | 4331 |
|
4332 |
lemmas word_split_bl_no_bin [simp] = |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4333 |
word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4334 |
|
65336 | 4335 |
text \<open> |
4336 |
This odd result arises from the fact that the statement of the |
|
4337 |
result implies that the decoded words are of the same type, |
|
4338 |
and therefore of the same length, as the original word.\<close> |
|
37660 | 4339 |
|
4340 |
lemma word_rsplit_same: "word_rsplit w = [w]" |
|
65336 | 4341 |
by (simp add: word_rsplit_def bin_rsplit_all) |
4342 |
||
4343 |
lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \<longleftrightarrow> size w = 0" |
|
4344 |
by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def |
|
4345 |
split: prod.split) |
|
37660 | 4346 |
|
4347 |
lemma test_bit_rsplit: |
|
65363 | 4348 |
"sw = word_rsplit w \<Longrightarrow> m < size (hd sw) \<Longrightarrow> |
4349 |
k < length sw \<Longrightarrow> (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" |
|
4350 |
for sw :: "'a::len word list" |
|
37660 | 4351 |
apply (unfold word_rsplit_def word_test_bit_def) |
4352 |
apply (rule trans) |
|
65336 | 4353 |
apply (rule_tac f = "\<lambda>x. bin_nth x m" in arg_cong) |
37660 | 4354 |
apply (rule nth_map [symmetric]) |
4355 |
apply simp |
|
4356 |
apply (rule bin_nth_rsplit) |
|
4357 |
apply simp_all |
|
4358 |
apply (simp add : word_size rev_map) |
|
4359 |
apply (rule trans) |
|
4360 |
defer |
|
4361 |
apply (rule map_ident [THEN fun_cong]) |
|
4362 |
apply (rule refl [THEN map_cong]) |
|
4363 |
apply (simp add : word_ubin.eq_norm) |
|
4364 |
apply (erule bin_rsplit_size_sign [OF len_gt_0 refl]) |
|
4365 |
done |
|
4366 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4367 |
lemma word_rcat_bl: "word_rcat wl = of_bl (concat (map to_bl wl))" |
65336 | 4368 |
by (auto simp: word_rcat_def to_bl_def' of_bl_def bin_rcat_bl) |
4369 |
||
4370 |
lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" |
|
4371 |
by (induct wl) (auto simp: word_size) |
|
37660 | 4372 |
|
4373 |
lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] |
|
4374 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4375 |
lemma nth_rcat_lem: |
70185 | 4376 |
"n < length (wl::'a word list) * LENGTH('a::len) \<Longrightarrow> |
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4377 |
rev (concat (map to_bl wl)) ! n = |
70185 | 4378 |
rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" |
65336 | 4379 |
apply (induct wl) |
37660 | 4380 |
apply clarsimp |
4381 |
apply (clarsimp simp add : nth_append size_rcat_lem) |
|
71997 | 4382 |
apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) |
4383 |
apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) |
|
37660 | 4384 |
done |
4385 |
||
4386 |
lemma test_bit_rcat: |
|
65363 | 4387 |
"sw = size (hd wl) \<Longrightarrow> rc = word_rcat wl \<Longrightarrow> rc !! n = |
65336 | 4388 |
(n < size rc \<and> n div sw < size wl \<and> (rev wl) ! (n div sw) !! (n mod sw))" |
65363 | 4389 |
for wl :: "'a::len word list" |
37660 | 4390 |
apply (unfold word_rcat_bl word_size) |
71997 | 4391 |
apply (clarsimp simp add: test_bit_of_bl size_rcat_lem word_size) |
4392 |
apply (metis div_le_mono len_gt_0 len_not_eq_0 less_mult_imp_div_less mod_less_divisor nonzero_mult_div_cancel_right not_le nth_rcat_lem test_bit_bl word_size) |
|
37660 | 4393 |
done |
4394 |
||
67399 | 4395 |
lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" |
65336 | 4396 |
for x :: "'a::comm_monoid_add" |
4397 |
by (induct xs arbitrary: x) (auto simp: add.assoc) |
|
37660 | 4398 |
|
4399 |
lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] |
|
4400 |
||
71996 | 4401 |
lemma test_bit_rsplit_alt: |
4402 |
\<open>(word_rsplit w :: 'b::len word list) ! i !! m \<longleftrightarrow> |
|
4403 |
w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\<close> |
|
4404 |
if \<open>i < length (word_rsplit w :: 'b::len word list)\<close> \<open>m < size (hd (word_rsplit w :: 'b::len word list))\<close> \<open>0 < length (word_rsplit w :: 'b::len word list)\<close> |
|
4405 |
for w :: \<open>'a::len word\<close> |
|
4406 |
apply (rule trans) |
|
4407 |
apply (rule test_bit_cong) |
|
71997 | 4408 |
apply (rule rev_nth [of _ \<open>rev (word_rsplit w)\<close>, simplified rev_rev_ident]) |
4409 |
apply simp |
|
71996 | 4410 |
apply (rule that(1)) |
71997 | 4411 |
apply simp |
71996 | 4412 |
apply (rule test_bit_rsplit) |
4413 |
apply (rule refl) |
|
4414 |
apply (rule asm_rl) |
|
4415 |
apply (rule that(2)) |
|
4416 |
apply (rule diff_Suc_less) |
|
4417 |
apply (rule that(3)) |
|
4418 |
done |
|
4419 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4420 |
lemma word_rsplit_len_indep [OF refl refl refl refl]: |
65268 | 4421 |
"[u,v] = p \<Longrightarrow> [su,sv] = q \<Longrightarrow> word_rsplit u = su \<Longrightarrow> |
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4422 |
word_rsplit v = sv \<Longrightarrow> length su = length sv" |
65336 | 4423 |
by (auto simp: word_rsplit_def bin_rsplit_len_indep) |
37660 | 4424 |
|
65268 | 4425 |
lemma length_word_rsplit_size: |
70185 | 4426 |
"n = LENGTH('a::len) \<Longrightarrow> |
65336 | 4427 |
length (word_rsplit w :: 'a word list) \<le> m \<longleftrightarrow> size w \<le> m * n" |
4428 |
by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) |
|
37660 | 4429 |
|
65268 | 4430 |
lemmas length_word_rsplit_lt_size = |
37660 | 4431 |
length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] |
4432 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4433 |
lemma length_word_rsplit_exp_size: |
70185 | 4434 |
"n = LENGTH('a::len) \<Longrightarrow> |
37660 | 4435 |
length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" |
65336 | 4436 |
by (auto simp: word_rsplit_def word_size bin_rsplit_len) |
37660 | 4437 |
|
65268 | 4438 |
lemma length_word_rsplit_even_size: |
70185 | 4439 |
"n = LENGTH('a::len) \<Longrightarrow> size w = m * n \<Longrightarrow> |
37660 | 4440 |
length (word_rsplit w :: 'a word list) = m" |
71997 | 4441 |
by (cases \<open>LENGTH('a)\<close>) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) |
37660 | 4442 |
|
4443 |
lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] |
|
4444 |
||
67408 | 4445 |
\<comment> \<open>alternative proof of \<open>word_rcat_rsplit\<close>\<close> |
66808
1907167b6038
elementary definition of division on natural numbers
haftmann
parents:
66453
diff
changeset
|
4446 |
lemmas tdle = times_div_less_eq_dividend |
71997 | 4447 |
lemmas dtle = xtrans(4) [OF tdle mult.commute] |
37660 | 4448 |
|
4449 |
lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" |
|
4450 |
apply (rule word_eqI) |
|
65336 | 4451 |
apply (clarsimp simp: test_bit_rcat word_size) |
37660 | 4452 |
apply (subst refl [THEN test_bit_rsplit]) |
65268 | 4453 |
apply (simp_all add: word_size |
37660 | 4454 |
refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) |
4455 |
apply safe |
|
71997 | 4456 |
apply (erule xtrans(7), rule dtle)+ |
37660 | 4457 |
done |
4458 |
||
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4459 |
lemma size_word_rsplit_rcat_size: |
70185 | 4460 |
"word_rcat ws = frcw \<Longrightarrow> size frcw = length ws * LENGTH('a) |
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4461 |
\<Longrightarrow> length (word_rsplit frcw::'a word list) = length ws" |
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
|
4462 |
for ws :: "'a::len word list" and frcw :: "'b::len word" |
71997 | 4463 |
by (cases \<open>LENGTH('a)\<close>) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) |
37660 | 4464 |
|
4465 |
lemma msrevs: |
|
65336 | 4466 |
"0 < n \<Longrightarrow> (k * n + m) div n = m div n + k" |
4467 |
"(k * n + m) mod n = m mod n" |
|
4468 |
for n :: nat |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57492
diff
changeset
|
4469 |
by (auto simp: add.commute) |
37660 | 4470 |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4471 |
lemma word_rsplit_rcat_size [OF refl]: |
65336 | 4472 |
"word_rcat ws = frcw \<Longrightarrow> |
70185 | 4473 |
size frcw = length ws * LENGTH('a) \<Longrightarrow> word_rsplit frcw = ws" |
65336 | 4474 |
for ws :: "'a::len word list" |
37660 | 4475 |
apply (frule size_word_rsplit_rcat_size, assumption) |
4476 |
apply (clarsimp simp add : word_size) |
|
4477 |
apply (rule nth_equalityI, assumption) |
|
4478 |
apply clarsimp |
|
46023
fad87bb608fc
restate some lemmas to respect int/bin distinction
huffman
parents:
46022
diff
changeset
|
4479 |
apply (rule word_eqI [rule_format]) |
37660 | 4480 |
apply (rule trans) |
4481 |
apply (rule test_bit_rsplit_alt) |
|
4482 |
apply (clarsimp simp: word_size)+ |
|
4483 |
apply (rule trans) |
|
65336 | 4484 |
apply (rule test_bit_rcat [OF refl refl]) |
55818 | 4485 |
apply (simp add: word_size) |
71997 | 4486 |
apply (subst rev_nth) |
37660 | 4487 |
apply arith |
71997 | 4488 |
apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) |
37660 | 4489 |
apply safe |
41550 | 4490 |
apply (simp add: diff_mult_distrib) |
65336 | 4491 |
apply (cases "size ws") |
4492 |
apply simp_all |
|
37660 | 4493 |
done |
4494 |
||
4495 |
||
61799 | 4496 |
subsection \<open>Rotation\<close> |
37660 | 4497 |
|
4498 |
lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] |
|
4499 |
||
71990 | 4500 |
lemma bit_word_rotl_iff: |
4501 |
\<open>bit (word_rotl m w) n \<longleftrightarrow> |
|
4502 |
n < LENGTH('a) \<and> bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\<close> |
|
4503 |
for w :: \<open>'a::len word\<close> |
|
4504 |
proof (cases \<open>n < LENGTH('a)\<close>) |
|
4505 |
case False |
|
4506 |
then show ?thesis |
|
4507 |
by (auto dest: bit_imp_le_length) |
|
4508 |
next |
|
4509 |
case True |
|
4510 |
define k where \<open>k = int n - int m\<close> |
|
4511 |
then have k: \<open>int n = k + int m\<close> |
|
4512 |
by simp |
|
4513 |
define l where \<open>l = int LENGTH('a)\<close> |
|
4514 |
then have l: \<open>int LENGTH('a) = l\<close> \<open>l > 0\<close> |
|
4515 |
by simp_all |
|
4516 |
have *: \<open>int (m - n) = int m - int n\<close> if \<open>n \<le> m\<close> for n m |
|
4517 |
using that by (simp add: int_minus) |
|
4518 |
from \<open>l > 0\<close> have \<open>l = 1 + (k mod l + ((- 1 - k) mod l))\<close> |
|
4519 |
using minus_mod_int_eq [of l \<open>k + 1\<close>] by (simp add: algebra_simps) |
|
4520 |
then have \<open>int (LENGTH('a) - Suc ((m + LENGTH('a) - Suc n) mod LENGTH('a))) = |
|
4521 |
int ((n + LENGTH('a) - m mod LENGTH('a)) mod LENGTH('a))\<close> |
|
4522 |
by (simp add: * k l zmod_int Suc_leI trans_le_add2 algebra_simps mod_simps \<open>n < LENGTH('a)\<close>) |
|
4523 |
then have \<open>LENGTH('a) - Suc ((m + LENGTH('a) - Suc n) mod LENGTH('a)) = |
|
4524 |
(n + LENGTH('a) - m mod LENGTH('a)) mod LENGTH('a)\<close> |
|
4525 |
by simp |
|
4526 |
with True show ?thesis |
|
4527 |
by (simp add: word_rotl_def bit_of_bl_iff rev_nth nth_rotate nth_to_bl) |
|
4528 |
qed |
|
4529 |
||
37660 | 4530 |
lemmas word_rot_defs = word_roti_def word_rotr_def word_rotl_def |
4531 |
||
65336 | 4532 |
lemma rotate_eq_mod: "m mod length xs = n mod length xs \<Longrightarrow> rotate m xs = rotate n xs" |
37660 | 4533 |
apply (rule box_equals) |
4534 |
defer |
|
4535 |
apply (rule rotate_conv_mod [symmetric])+ |
|
4536 |
apply simp |
|
4537 |
done |
|
4538 |
||
65268 | 4539 |
lemmas rotate_eqs = |
37660 | 4540 |
trans [OF rotate0 [THEN fun_cong] id_apply] |
65268 | 4541 |
rotate_rotate [symmetric] |
45604 | 4542 |
rotate_id |
65268 | 4543 |
rotate_conv_mod |
37660 | 4544 |
rotate_eq_mod |
4545 |
||
4546 |
||
61799 | 4547 |
subsubsection \<open>Rotation of list to right\<close> |
37660 | 4548 |
|
4549 |
lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" |
|
65336 | 4550 |
by (cases l) (auto simp: rotater1_def) |
37660 | 4551 |
|
4552 |
lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" |
|
4553 |
apply (unfold rotater1_def) |
|
4554 |
apply (cases "l") |
|
65336 | 4555 |
apply (case_tac [2] "list") |
4556 |
apply auto |
|
37660 | 4557 |
done |
4558 |
||
4559 |
lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" |
|
65336 | 4560 |
by (cases l) (auto simp: rotater1_def) |
37660 | 4561 |
|
4562 |
lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" |
|
65336 | 4563 |
by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') |
65268 | 4564 |
|
37660 | 4565 |
lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" |
65336 | 4566 |
by (induct n) (auto simp: rotater_def intro: rotater1_rev') |
37660 | 4567 |
|
45816
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4568 |
lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" |
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
huffman
parents:
45811
diff
changeset
|
4569 |
using rotater_rev' [where xs = "rev ys"] by simp |
37660 | 4570 |
|
65268 | 4571 |
lemma rotater_drop_take: |
4572 |
"rotater n xs = |
|
65336 | 4573 |
drop (length xs - n mod length xs) xs @ |
4574 |
take (length xs - n mod length xs) xs" |
|
4575 |
by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) |
|
4576 |
||
4577 |
lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" |
|
37660 | 4578 |
unfolding rotater_def by auto |
4579 |
||
71990 | 4580 |
lemma nth_rotater: |
4581 |
\<open>rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\<close> if \<open>n < length xs\<close> |
|
4582 |
using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) |
|
4583 |
||
4584 |
lemma nth_rotater1: |
|
4585 |
\<open>rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\<close> if \<open>n < length xs\<close> |
|
4586 |
using that nth_rotater [of n xs 1] by simp |
|
4587 |
||
4588 |
lemma rotate_inv_plus [rule_format]: |
|
65336 | 4589 |
"\<forall>k. k = m + n \<longrightarrow> rotater k (rotate n xs) = rotater m xs \<and> |
4590 |
rotate k (rotater n xs) = rotate m xs \<and> |
|
4591 |
rotater n (rotate k xs) = rotate m xs \<and> |
|
37660 | 4592 |
rotate n (rotater k xs) = rotater m xs" |
65336 | 4593 |
by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) |
65268 | 4594 |
|
37660 | 4595 |
lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] |
4596 |
||
4597 |
lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] |
|
4598 |
||
45604 | 4599 |
lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] |
4600 |
lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] |
|
37660 | 4601 |
|
65336 | 4602 |
lemma rotate_gal: "rotater n xs = ys \<longleftrightarrow> rotate n ys = xs" |
37660 | 4603 |
by auto |
4604 |
||
65336 | 4605 |
lemma rotate_gal': "ys = rotater n xs \<longleftrightarrow> xs = rotate n ys" |
37660 | 4606 |
by auto |
4607 |
||
65336 | 4608 |
lemma length_rotater [simp]: "length (rotater n xs) = length xs" |
37660 | 4609 |
by (simp add : rotater_rev) |
4610 |
||
71990 | 4611 |
lemma bit_word_rotr_iff: |
4612 |
\<open>bit (word_rotr m w) n \<longleftrightarrow> |
|
4613 |
n < LENGTH('a) \<and> bit w ((n + m) mod LENGTH('a))\<close> |
|
4614 |
for w :: \<open>'a::len word\<close> |
|
4615 |
proof (cases \<open>n < LENGTH('a)\<close>) |
|
4616 |
case False |
|
4617 |
then show ?thesis |
|
4618 |
by (auto dest: bit_imp_le_length) |
|
4619 |
next |
|
4620 |
case True |
|
4621 |
define k where \<open>k = int n + int m\<close> |
|
4622 |
then have k: \<open>int n = k - int m\<close> |
|
4623 |
by simp |
|
4624 |
define l where \<open>l = int LENGTH('a)\<close> |
|
4625 |
then have l: \<open>int LENGTH('a) = l\<close> \<open>l > 0\<close> |
|
4626 |
by simp_all |
|
4627 |
have *: \<open>int (m - n) = int m - int n\<close> if \<open>n \<le> m\<close> for n m |
|
4628 |
using that by (simp add: int_minus) |
|
4629 |
have \<open>int ((LENGTH('a) |
|
4630 |
- Suc ((LENGTH('a) + LENGTH('a) - Suc (n + m mod LENGTH('a))) mod LENGTH('a)))) = |
|
4631 |
int ((n + m) mod LENGTH('a))\<close> |
|
4632 |
using True |
|
4633 |
apply (simp add: l * zmod_int Suc_leI add_strict_mono) |
|
4634 |
apply (subst mod_diff_left_eq [symmetric]) |
|
4635 |
apply simp |
|
4636 |
using l minus_mod_int_eq [of l \<open>int n + int m mod l + 1\<close>] apply simp |
|
4637 |
apply (simp add: mod_simps) |
|
4638 |
done |
|
4639 |
then have \<open>(LENGTH('a) |
|
4640 |
- Suc ((LENGTH('a) + LENGTH('a) - Suc (n + m mod LENGTH('a))) mod LENGTH('a))) = |
|
4641 |
((n + m) mod LENGTH('a))\<close> |
|
4642 |
by simp |
|
4643 |
with True show ?thesis |
|
4644 |
by (simp add: word_rotr_def bit_of_bl_iff rev_nth nth_rotater nth_to_bl) |
|
4645 |
qed |
|
4646 |
||
4647 |
lemma bit_word_roti_iff: |
|
4648 |
\<open>bit (word_roti k w) n \<longleftrightarrow> |
|
4649 |
n < LENGTH('a) \<and> bit w (nat ((int n + k) mod int LENGTH('a)))\<close> |
|
4650 |
for w :: \<open>'a::len word\<close> |
|
4651 |
proof (cases \<open>k \<ge> 0\<close>) |
|
4652 |
case True |
|
4653 |
moreover define m where \<open>m = nat k\<close> |
|
4654 |
ultimately have \<open>k = int m\<close> |
|
4655 |
by simp |
|
4656 |
then show ?thesis |
|
4657 |
by (simp add: word_roti_def bit_word_rotr_iff nat_mod_distrib nat_add_distrib) |
|
4658 |
next |
|
4659 |
have *: \<open>int (m - n) = int m - int n\<close> if \<open>n \<le> m\<close> for n m |
|
4660 |
using that by (simp add: int_minus) |
|
4661 |
case False |
|
4662 |
moreover define m where \<open>m = nat (- k)\<close> |
|
4663 |
ultimately have \<open>k = - int m\<close> \<open>k < 0\<close> |
|
4664 |
by simp_all |
|
4665 |
moreover have \<open>(int n - int m) mod int LENGTH('a) = |
|
4666 |
int ((n + LENGTH('a) - m mod LENGTH('a)) mod LENGTH('a))\<close> |
|
4667 |
apply (simp add: zmod_int * trans_le_add2 mod_simps) |
|
4668 |
apply (metis mod_add_self2 mod_diff_cong) |
|
4669 |
done |
|
4670 |
ultimately show ?thesis |
|
4671 |
by (simp add: word_roti_def bit_word_rotl_iff nat_mod_distrib) |
|
4672 |
qed |
|
4673 |
||
65336 | 4674 |
lemma restrict_to_left: "x = y \<Longrightarrow> x = z \<longleftrightarrow> y = z" |
4675 |
by simp |
|
38527 | 4676 |
|
65268 | 4677 |
lemmas rrs0 = rotate_eqs [THEN restrict_to_left, |
45604 | 4678 |
simplified rotate_gal [symmetric] rotate_gal' [symmetric]] |
37660 | 4679 |
lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] |
45604 | 4680 |
lemmas rotater_eqs = rrs1 [simplified length_rotater] |
37660 | 4681 |
lemmas rotater_0 = rotater_eqs (1) |
4682 |
lemmas rotater_add = rotater_eqs (2) |
|
4683 |
||
4684 |
||
61799 | 4685 |
subsubsection \<open>map, map2, commuting with rotate(r)\<close> |
37660 | 4686 |
|
65336 | 4687 |
lemma butlast_map: "xs \<noteq> [] \<Longrightarrow> butlast (map f xs) = map f (butlast xs)" |
37660 | 4688 |
by (induct xs) auto |
4689 |
||
65268 | 4690 |
lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" |
65336 | 4691 |
by (cases xs) (auto simp: rotater1_def last_map butlast_map) |
4692 |
||
4693 |
lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" |
|
4694 |
by (induct n) (auto simp: rotater_def rotater1_map) |
|
37660 | 4695 |
|
4696 |
lemma but_last_zip [rule_format] : |
|
65336 | 4697 |
"\<forall>ys. length xs = length ys \<longrightarrow> xs \<noteq> [] \<longrightarrow> |
4698 |
last (zip xs ys) = (last xs, last ys) \<and> |
|
65268 | 4699 |
butlast (zip xs ys) = zip (butlast xs) (butlast ys)" |
65336 | 4700 |
apply (induct xs) |
4701 |
apply auto |
|
37660 | 4702 |
apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ |
4703 |
done |
|
4704 |
||
4705 |
lemma but_last_map2 [rule_format] : |
|
65336 | 4706 |
"\<forall>ys. length xs = length ys \<longrightarrow> xs \<noteq> [] \<longrightarrow> |
4707 |
last (map2 f xs ys) = f (last xs) (last ys) \<and> |
|
65268 | 4708 |
butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" |
65336 | 4709 |
apply (induct xs) |
4710 |
apply auto |
|
37660 | 4711 |
apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ |
4712 |
done |
|
4713 |
||
4714 |
lemma rotater1_zip: |
|
65268 | 4715 |
"length xs = length ys \<Longrightarrow> |
4716 |
rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" |
|
37660 | 4717 |
apply (unfold rotater1_def) |
65336 | 4718 |
apply (cases xs) |
37660 | 4719 |
apply auto |
4720 |
apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ |
|
4721 |
done |
|
4722 |
||
4723 |
lemma rotater1_map2: |
|
65268 | 4724 |
"length xs = length ys \<Longrightarrow> |
4725 |
rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" |
|
70193 | 4726 |
by (simp add: rotater1_map rotater1_zip) |
37660 | 4727 |
|
65268 | 4728 |
lemmas lrth = |
4729 |
box_equals [OF asm_rl length_rotater [symmetric] |
|
4730 |
length_rotater [symmetric], |
|
37660 | 4731 |
THEN rotater1_map2] |
4732 |
||
65268 | 4733 |
lemma rotater_map2: |
4734 |
"length xs = length ys \<Longrightarrow> |
|
4735 |
rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" |
|
37660 | 4736 |
by (induct n) (auto intro!: lrth) |
4737 |
||
4738 |
lemma rotate1_map2: |
|
65268 | 4739 |
"length xs = length ys \<Longrightarrow> |
4740 |
rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" |
|
70193 | 4741 |
by (cases xs; cases ys) auto |
37660 | 4742 |
|
65268 | 4743 |
lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] |
37660 | 4744 |
length_rotate [symmetric], THEN rotate1_map2] |
4745 |
||
65268 | 4746 |
lemma rotate_map2: |
4747 |
"length xs = length ys \<Longrightarrow> |
|
4748 |
rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" |
|
37660 | 4749 |
by (induct n) (auto intro!: lth) |
4750 |
||
4751 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67408
diff
changeset
|
4752 |
\<comment> \<open>corresponding equalities for word rotation\<close> |
37660 | 4753 |
|
65336 | 4754 |
lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" |
37660 | 4755 |
by (simp add: word_bl.Abs_inverse' word_rotl_def) |
4756 |
||
4757 |
lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] |
|
4758 |
||
4759 |
lemmas word_rotl_eqs = |
|
45538
1fffa81b9b83
eliminated slightly odd Rep' with dynamically-scoped [simplified];
wenzelm
parents:
45529
diff
changeset
|
4760 |
blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] |
37660 | 4761 |
|
65336 | 4762 |
lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" |
37660 | 4763 |
by (simp add: word_bl.Abs_inverse' word_rotr_def) |
4764 |
||
4765 |
lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] |
|
4766 |
||
4767 |
lemmas word_rotr_eqs = |
|
45538
1fffa81b9b83
eliminated slightly odd Rep' with dynamically-scoped [simplified];
wenzelm
parents:
45529
diff
changeset
|
4768 |
brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] |
37660 | 4769 |
|
4770 |
declare word_rotr_eqs (1) [simp] |
|
4771 |
declare word_rotl_eqs (1) [simp] |
|
4772 |
||
65336 | 4773 |
lemma word_rot_rl [simp]: "word_rotl k (word_rotr k v) = v" |
4774 |
and word_rot_lr [simp]: "word_rotr k (word_rotl k v) = v" |
|
37660 | 4775 |
by (auto simp add: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric]) |
4776 |
||
65336 | 4777 |
lemma word_rot_gal: "word_rotr n v = w \<longleftrightarrow> word_rotl n w = v" |
4778 |
and word_rot_gal': "w = word_rotr n v \<longleftrightarrow> v = word_rotl n w" |
|
4779 |
by (auto simp: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric] dest: sym) |
|
4780 |
||
4781 |
lemma word_rotr_rev: "word_rotr n w = word_reverse (word_rotl n (word_reverse w))" |
|
4782 |
by (simp only: word_bl.Rep_inject [symmetric] to_bl_word_rev to_bl_rotr to_bl_rotl rotater_rev) |
|
65268 | 4783 |
|
37660 | 4784 |
lemma word_roti_0 [simp]: "word_roti 0 w = w" |
65336 | 4785 |
by (auto simp: word_rot_defs) |
37660 | 4786 |
|
4787 |
lemmas abl_cong = arg_cong [where f = "of_bl"] |
|
4788 |
||
65336 | 4789 |
lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" |
37660 | 4790 |
proof - |
65336 | 4791 |
have rotater_eq_lem: "\<And>m n xs. m = n \<Longrightarrow> rotater m xs = rotater n xs" |
37660 | 4792 |
by auto |
4793 |
||
65336 | 4794 |
have rotate_eq_lem: "\<And>m n xs. m = n \<Longrightarrow> rotate m xs = rotate n xs" |
37660 | 4795 |
by auto |
4796 |
||
65268 | 4797 |
note rpts [symmetric] = |
37660 | 4798 |
rotate_inv_plus [THEN conjunct1] |
4799 |
rotate_inv_plus [THEN conjunct2, THEN conjunct1] |
|
4800 |
rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct1] |
|
4801 |
rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct2] |
|
4802 |
||
4803 |
note rrp = trans [symmetric, OF rotate_rotate rotate_eq_lem] |
|
4804 |
note rrrp = trans [symmetric, OF rotater_add [symmetric] rotater_eq_lem] |
|
4805 |
||
4806 |
show ?thesis |
|
65336 | 4807 |
apply (unfold word_rot_defs) |
4808 |
apply (simp only: split: if_split) |
|
4809 |
apply (safe intro!: abl_cong) |
|
4810 |
apply (simp_all only: to_bl_rotl [THEN word_bl.Rep_inverse'] |
|
4811 |
to_bl_rotl |
|
4812 |
to_bl_rotr [THEN word_bl.Rep_inverse'] |
|
4813 |
to_bl_rotr) |
|
4814 |
apply (rule rrp rrrp rpts, |
|
4815 |
simp add: nat_add_distrib [symmetric] |
|
4816 |
nat_diff_distrib [symmetric])+ |
|
4817 |
done |
|
37660 | 4818 |
qed |
65268 | 4819 |
|
67118 | 4820 |
lemma word_roti_conv_mod': |
4821 |
"word_roti n w = word_roti (n mod int (size w)) w" |
|
4822 |
proof (cases "size w = 0") |
|
4823 |
case True |
|
4824 |
then show ?thesis |
|
4825 |
by simp |
|
4826 |
next |
|
4827 |
case False |
|
4828 |
then have [simp]: "l mod int (size w) \<ge> 0" for l |
|
4829 |
by simp |
|
4830 |
then have *: "word_roti (n mod int (size w)) w = word_rotr (nat (n mod int (size w))) w" |
|
4831 |
by (simp add: word_roti_def) |
|
4832 |
show ?thesis |
|
4833 |
proof (cases "n \<ge> 0") |
|
4834 |
case True |
|
4835 |
then show ?thesis |
|
4836 |
apply (auto simp add: not_le *) |
|
4837 |
apply (auto simp add: word_rot_defs) |
|
4838 |
apply (safe intro!: abl_cong) |
|
4839 |
apply (rule rotater_eqs) |
|
4840 |
apply (simp add: word_size nat_mod_distrib) |
|
4841 |
done |
|
4842 |
next |
|
4843 |
case False |
|
4844 |
moreover define k where "k = - n" |
|
4845 |
ultimately have n: "n = - k" |
|
4846 |
by simp_all |
|
4847 |
from False show ?thesis |
|
4848 |
apply (auto simp add: not_le *) |
|
4849 |
apply (auto simp add: word_rot_defs) |
|
4850 |
apply (simp add: n) |
|
4851 |
apply (safe intro!: abl_cong) |
|
4852 |
apply (simp add: rotater_add [symmetric] rotate_gal [symmetric]) |
|
4853 |
apply (rule rotater_eqs) |
|
4854 |
apply (simp add: word_size [symmetric, of w]) |
|
4855 |
apply (rule of_nat_eq_0_iff [THEN iffD1]) |
|
4856 |
apply (auto simp add: nat_add_distrib [symmetric] mod_eq_0_iff_dvd) |
|
4857 |
using dvd_nat_abs_iff [of "size w" "- k mod int (size w) + k"] |
|
4858 |
apply (simp add: algebra_simps) |
|
4859 |
apply (simp add: word_size) |
|
71942 | 4860 |
apply (metis dvd_eq_mod_eq_0 eq_neg_iff_add_eq_0 k_def mod_0 mod_add_right_eq n) |
67118 | 4861 |
done |
4862 |
qed |
|
4863 |
qed |
|
37660 | 4864 |
|
4865 |
lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] |
|
4866 |
||
4867 |
||
61799 | 4868 |
subsubsection \<open>"Word rotation commutes with bit-wise operations\<close> |
37660 | 4869 |
|
67408 | 4870 |
\<comment> \<open>using locale to not pollute lemma namespace\<close> |
65268 | 4871 |
locale word_rotate |
37660 | 4872 |
begin |
4873 |
||
4874 |
lemmas word_rot_defs' = to_bl_rotl to_bl_rotr |
|
4875 |
||
4876 |
lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor |
|
4877 |
||
45538
1fffa81b9b83
eliminated slightly odd Rep' with dynamically-scoped [simplified];
wenzelm
parents:
45529
diff
changeset
|
4878 |
lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] |
37660 | 4879 |
|
4880 |
lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 |
|
4881 |
||
45604 | 4882 |
lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v |
37660 | 4883 |
|
4884 |
lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map |
|
4885 |
||
4886 |
lemma word_rot_logs: |
|
71149 | 4887 |
"word_rotl n (NOT v) = NOT (word_rotl n v)" |
4888 |
"word_rotr n (NOT v) = NOT (word_rotr n v)" |
|
37660 | 4889 |
"word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" |
4890 |
"word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" |
|
4891 |
"word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" |
|
4892 |
"word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" |
|
4893 |
"word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" |
|
65268 | 4894 |
"word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" |
37660 | 4895 |
by (rule word_bl.Rep_eqD, |
4896 |
rule word_rot_defs' [THEN trans], |
|
4897 |
simp only: blwl_syms [symmetric], |
|
65268 | 4898 |
rule th1s [THEN trans], |
37660 | 4899 |
rule refl)+ |
4900 |
end |
|
4901 |
||
4902 |
lemmas word_rot_logs = word_rotate.word_rot_logs |
|
4903 |
||
4904 |
lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, |
|
45604 | 4905 |
simplified word_bl_Rep'] |
37660 | 4906 |
|
4907 |
lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, |
|
45604 | 4908 |
simplified word_bl_Rep'] |
37660 | 4909 |
|
65268 | 4910 |
lemma bl_word_roti_dt': |
4911 |
"n = nat ((- i) mod int (size (w :: 'a::len word))) \<Longrightarrow> |
|
37660 | 4912 |
to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" |
4913 |
apply (unfold word_roti_def) |
|
4914 |
apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) |
|
4915 |
apply safe |
|
4916 |
apply (simp add: zmod_zminus1_eq_if) |
|
4917 |
apply safe |
|
4918 |
apply (simp add: nat_mult_distrib) |
|
65268 | 4919 |
apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj |
37660 | 4920 |
[THEN conjunct2, THEN order_less_imp_le]] |
4921 |
nat_mod_distrib) |
|
4922 |
apply (simp add: nat_mod_distrib) |
|
4923 |
done |
|
4924 |
||
4925 |
lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] |
|
4926 |
||
65268 | 4927 |
lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] |
45604 | 4928 |
lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] |
4929 |
lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] |
|
37660 | 4930 |
|
65336 | 4931 |
lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \<and> word_rotl i 0 = 0" |
4932 |
by (simp add: word_rotr_dt word_rotl_dt replicate_add [symmetric]) |
|
37660 | 4933 |
|
4934 |
lemma word_roti_0' [simp] : "word_roti n 0 = 0" |
|
65336 | 4935 |
by (auto simp: word_roti_def) |
37660 | 4936 |
|
65268 | 4937 |
lemmas word_rotr_dt_no_bin' [simp] = |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4938 |
word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4939 |
(* FIXME: negative numerals, 0 and 1 *) |
37660 | 4940 |
|
65268 | 4941 |
lemmas word_rotl_dt_no_bin' [simp] = |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4942 |
word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4943 |
(* FIXME: negative numerals, 0 and 1 *) |
37660 | 4944 |
|
4945 |
declare word_roti_def [simp] |
|
4946 |
||
4947 |
||
61799 | 4948 |
subsection \<open>Maximum machine word\<close> |
37660 | 4949 |
|
4950 |
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
|
4951 |
fixes x :: "'a::len word" |
70185 | 4952 |
obtains n where "x = word_of_int n" and "0 \<le> n" and "n < 2^LENGTH('a)" |
37660 | 4953 |
by (cases x rule: word_uint.Abs_cases) (simp add: uints_num) |
4954 |
||
4955 |
lemma word_nat_cases [cases type: word]: |
|
65336 | 4956 |
fixes x :: "'a::len word" |
70185 | 4957 |
obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" |
37660 | 4958 |
by (cases x rule: word_unat.Abs_cases) (simp add: unats_def) |
4959 |
||
71946 | 4960 |
lemma max_word_max [intro!]: "n \<le> max_word" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
4961 |
by (fact word_order.extremum) |
65268 | 4962 |
|
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
|
4963 |
lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
4964 |
by (subst word_uint.Abs_norm [symmetric]) simp |
37660 | 4965 |
|
70185 | 4966 |
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
|
4967 |
by (fact word_exp_length_eq_0) |
37660 | 4968 |
|
4969 |
lemma max_word_wrap: "x + 1 = 0 \<Longrightarrow> x = max_word" |
|
71946 | 4970 |
by (simp add: eq_neg_iff_add_eq_0) |
4971 |
||
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
|
4972 |
lemma max_word_bl: "to_bl (max_word::'a::len word) = replicate LENGTH('a) True" |
71946 | 4973 |
by (fact to_bl_n1) |
4974 |
||
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
|
4975 |
lemma max_test_bit: "(max_word::'a::len word) !! n \<longleftrightarrow> n < LENGTH('a)" |
71946 | 4976 |
by (fact nth_minus1) |
4977 |
||
4978 |
lemma word_and_max: "x AND max_word = x" |
|
4979 |
by (fact word_log_esimps) |
|
4980 |
||
4981 |
lemma word_or_max: "x OR max_word = max_word" |
|
4982 |
by (fact word_log_esimps) |
|
37660 | 4983 |
|
65336 | 4984 |
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
|
4985 |
for x y z :: "'a::len word" |
37660 | 4986 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
4987 |
||
65336 | 4988 |
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
|
4989 |
for x y z :: "'a::len word" |
37660 | 4990 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
4991 |
||
65336 | 4992 |
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
|
4993 |
for x :: "'a::len word" |
37660 | 4994 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
4995 |
||
65336 | 4996 |
lemma word_or_not [simp]: "x OR NOT x = max_word" |
37660 | 4997 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
4998 |
||
65336 | 4999 |
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
|
5000 |
for x y :: "'a::len word" |
37660 | 5001 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
5002 |
||
65336 | 5003 |
lemma shiftr_x_0 [iff]: "x >> 0 = 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
|
5004 |
for x :: "'a::len word" |
37660 | 5005 |
by (simp add: shiftr_bl) |
5006 |
||
65336 | 5007 |
lemma shiftl_x_0 [simp]: "x << 0 = x" |
5008 |
for x :: "'a::len word" |
|
37660 | 5009 |
by (simp add: shiftl_t2n) |
5010 |
||
65336 | 5011 |
lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" |
37660 | 5012 |
by (simp add: shiftl_t2n) |
5013 |
||
65336 | 5014 |
lemma uint_lt_0 [simp]: "uint x < 0 = False" |
37660 | 5015 |
by (simp add: linorder_not_less) |
5016 |
||
65336 | 5017 |
lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" |
45995
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
huffman
parents:
45958
diff
changeset
|
5018 |
unfolding shiftr1_def by simp |
37660 | 5019 |
|
65336 | 5020 |
lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" |
37660 | 5021 |
by (induct n) (auto simp: shiftr_def) |
5022 |
||
65336 | 5023 |
lemma word_less_1 [simp]: "x < 1 \<longleftrightarrow> x = 0" |
5024 |
for x :: "'a::len word" |
|
37660 | 5025 |
by (simp add: word_less_nat_alt unat_0_iff) |
5026 |
||
5027 |
lemma to_bl_mask: |
|
65268 | 5028 |
"to_bl (mask n :: 'a::len word) = |
70185 | 5029 |
replicate (LENGTH('a) - n) False @ |
5030 |
replicate (min (LENGTH('a)) n) True" |
|
37660 | 5031 |
by (simp add: mask_bl word_rep_drop min_def) |
5032 |
||
5033 |
lemma map_replicate_True: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
5034 |
"n = length xs \<Longrightarrow> |
65336 | 5035 |
map (\<lambda>(x,y). x \<and> y) (zip xs (replicate n True)) = xs" |
37660 | 5036 |
by (induct xs arbitrary: n) auto |
5037 |
||
5038 |
lemma map_replicate_False: |
|
65336 | 5039 |
"n = length xs \<Longrightarrow> map (\<lambda>(x,y). x \<and> y) |
37660 | 5040 |
(zip xs (replicate n False)) = replicate n False" |
5041 |
by (induct xs arbitrary: n) auto |
|
5042 |
||
5043 |
lemma bl_and_mask: |
|
5044 |
fixes w :: "'a::len word" |
|
65336 | 5045 |
and n :: nat |
70185 | 5046 |
defines "n' \<equiv> LENGTH('a) - n" |
65336 | 5047 |
shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" |
65268 | 5048 |
proof - |
37660 | 5049 |
note [simp] = map_replicate_True map_replicate_False |
67399 | 5050 |
have "to_bl (w AND mask n) = map2 (\<and>) (to_bl w) (to_bl (mask n::'a::len word))" |
37660 | 5051 |
by (simp add: bl_word_and) |
65336 | 5052 |
also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" |
5053 |
by simp |
|
67399 | 5054 |
also have "map2 (\<and>) \<dots> (to_bl (mask n::'a::len word)) = |
65336 | 5055 |
replicate n' False @ drop n' (to_bl w)" |
70193 | 5056 |
unfolding to_bl_mask n'_def by (subst zip_append) auto |
65336 | 5057 |
finally show ?thesis . |
37660 | 5058 |
qed |
5059 |
||
5060 |
lemma drop_rev_takefill: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
5061 |
"length xs \<le> n \<Longrightarrow> |
37660 | 5062 |
drop (n - length xs) (rev (takefill False n (rev xs))) = xs" |
5063 |
by (simp add: takefill_alt rev_take) |
|
5064 |
||
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
|
5065 |
lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" |
37660 | 5066 |
by (induct xs) auto |
5067 |
||
5068 |
lemma uint_plus_if_size: |
|
65268 | 5069 |
"uint (x + y) = |
65336 | 5070 |
(if uint x + uint y < 2^size x |
5071 |
then uint x + uint y |
|
5072 |
else uint x + uint y - 2^size x)" |
|
5073 |
by (simp add: word_arith_wis int_word_uint mod_add_if_z word_size) |
|
37660 | 5074 |
|
5075 |
lemma unat_plus_if_size: |
|
65363 | 5076 |
"unat (x + y) = |
65336 | 5077 |
(if unat x + unat y < 2^size x |
5078 |
then unat x + unat y |
|
5079 |
else unat x + unat y - 2^size x)" |
|
65363 | 5080 |
for x y :: "'a::len word" |
37660 | 5081 |
apply (subst word_arith_nat_defs) |
5082 |
apply (subst unat_of_nat) |
|
71997 | 5083 |
apply (auto simp add: not_less word_size) |
5084 |
apply (metis not_le unat_plus_if' unat_word_ariths(1)) |
|
37660 | 5085 |
done |
5086 |
||
65336 | 5087 |
lemma word_neq_0_conv: "w \<noteq> 0 \<longleftrightarrow> 0 < w" |
5088 |
for w :: "'a::len word" |
|
5089 |
by (simp add: word_gt_0) |
|
5090 |
||
5091 |
lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" |
|
5092 |
for c :: "'a::len word" |
|
55818 | 5093 |
by (fact unat_div) |
37660 | 5094 |
|
5095 |
lemma uint_sub_if_size: |
|
65268 | 5096 |
"uint (x - y) = |
65336 | 5097 |
(if uint y \<le> uint x |
5098 |
then uint x - uint y |
|
5099 |
else uint x - uint y + 2^size x)" |
|
5100 |
by (simp add: word_arith_wis int_word_uint mod_sub_if_z word_size) |
|
5101 |
||
5102 |
lemma unat_sub: "b \<le> a \<Longrightarrow> unat (a - b) = unat a - unat b" |
|
37660 | 5103 |
by (simp add: unat_def uint_sub_if_size word_le_def nat_diff_distrib) |
5104 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
5105 |
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
|
5106 |
lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w |
65268 | 5107 |
|
70185 | 5108 |
lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" |
37660 | 5109 |
proof - |
70185 | 5110 |
have *: "2^LENGTH('a) - i = -i + 2^LENGTH('a)" |
65336 | 5111 |
by simp |
37660 | 5112 |
show ?thesis |
65336 | 5113 |
apply (subst *) |
37660 | 5114 |
apply (subst word_uint.Abs_norm [symmetric], subst mod_add_self2) |
5115 |
apply simp |
|
5116 |
done |
|
5117 |
qed |
|
65268 | 5118 |
|
5119 |
lemmas word_of_int_inj = |
|
37660 | 5120 |
word_uint.Abs_inject [unfolded uints_num, simplified] |
5121 |
||
65336 | 5122 |
lemma word_le_less_eq: "x \<le> y \<longleftrightarrow> x = y \<or> x < y" |
5123 |
for x y :: "'z::len word" |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
5124 |
by (auto simp add: order_class.le_less) |
37660 | 5125 |
|
5126 |
lemma mod_plus_cong: |
|
65336 | 5127 |
fixes b b' :: int |
5128 |
assumes 1: "b = b'" |
|
5129 |
and 2: "x mod b' = x' mod b'" |
|
5130 |
and 3: "y mod b' = y' mod b'" |
|
5131 |
and 4: "x' + y' = z'" |
|
37660 | 5132 |
shows "(x + y) mod b = z' mod b'" |
5133 |
proof - |
|
5134 |
from 1 2[symmetric] 3[symmetric] 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
|
5135 |
by (simp add: mod_add_eq) |
37660 | 5136 |
also have "\<dots> = (x' + y') mod b'" |
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
64243
diff
changeset
|
5137 |
by (simp add: mod_add_eq) |
65336 | 5138 |
finally show ?thesis |
5139 |
by (simp add: 4) |
|
37660 | 5140 |
qed |
5141 |
||
5142 |
lemma mod_minus_cong: |
|
65336 | 5143 |
fixes b b' :: int |
5144 |
assumes "b = b'" |
|
5145 |
and "x mod b' = x' mod b'" |
|
5146 |
and "y mod b' = y' mod b'" |
|
5147 |
and "x' - y' = z'" |
|
37660 | 5148 |
shows "(x - y) mod b = z' mod b'" |
65336 | 5149 |
using assms [symmetric] by (auto intro: mod_diff_cong) |
5150 |
||
65363 | 5151 |
lemma word_induct_less: "P 0 \<Longrightarrow> (\<And>n. n < m \<Longrightarrow> P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P m" |
65336 | 5152 |
for P :: "'a::len word \<Rightarrow> bool" |
37660 | 5153 |
apply (cases m) |
5154 |
apply atomize |
|
5155 |
apply (erule rev_mp)+ |
|
5156 |
apply (rule_tac x=m in spec) |
|
5157 |
apply (induct_tac n) |
|
5158 |
apply simp |
|
5159 |
apply clarsimp |
|
5160 |
apply (erule impE) |
|
5161 |
apply clarsimp |
|
5162 |
apply (erule_tac x=n in allE) |
|
5163 |
apply (erule impE) |
|
5164 |
apply (simp add: unat_arith_simps) |
|
5165 |
apply (clarsimp simp: unat_of_nat) |
|
5166 |
apply simp |
|
5167 |
apply (erule_tac x="of_nat na" in allE) |
|
5168 |
apply (erule impE) |
|
5169 |
apply (simp add: unat_arith_simps) |
|
5170 |
apply (clarsimp simp: unat_of_nat) |
|
5171 |
apply simp |
|
5172 |
done |
|
65268 | 5173 |
|
65363 | 5174 |
lemma word_induct: "P 0 \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P m" |
65336 | 5175 |
for P :: "'a::len word \<Rightarrow> bool" |
5176 |
by (erule word_induct_less) simp |
|
5177 |
||
65363 | 5178 |
lemma word_induct2 [induct type]: "P 0 \<Longrightarrow> (\<And>n. 1 + n \<noteq> 0 \<Longrightarrow> P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P n" |
65336 | 5179 |
for P :: "'b::len word \<Rightarrow> bool" |
5180 |
apply (rule word_induct) |
|
5181 |
apply simp |
|
5182 |
apply (case_tac "1 + n = 0") |
|
5183 |
apply auto |
|
37660 | 5184 |
done |
5185 |
||
55816
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
haftmann
parents:
55415
diff
changeset
|
5186 |
|
61799 | 5187 |
subsection \<open>Recursion combinator for words\<close> |
46010 | 5188 |
|
54848 | 5189 |
definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a" |
65336 | 5190 |
where "word_rec forZero forSuc n = rec_nat forZero (forSuc \<circ> of_nat) (unat n)" |
37660 | 5191 |
|
5192 |
lemma word_rec_0: "word_rec z s 0 = z" |
|
5193 |
by (simp add: word_rec_def) |
|
5194 |
||
65363 | 5195 |
lemma word_rec_Suc: "1 + n \<noteq> 0 \<Longrightarrow> word_rec z s (1 + n) = s n (word_rec z s n)" |
5196 |
for n :: "'a::len word" |
|
71997 | 5197 |
apply (auto simp add: word_rec_def unat_word_ariths) |
5198 |
apply (metis (mono_tags, lifting) old.nat.simps(7) unatSuc word_unat.Rep_inverse word_unat.eq_norm word_unat.td_th) |
|
37660 | 5199 |
done |
5200 |
||
65363 | 5201 |
lemma word_rec_Pred: "n \<noteq> 0 \<Longrightarrow> word_rec z s n = s (n - 1) (word_rec z s (n - 1))" |
37660 | 5202 |
apply (rule subst[where t="n" and s="1 + (n - 1)"]) |
5203 |
apply simp |
|
5204 |
apply (subst word_rec_Suc) |
|
5205 |
apply simp |
|
5206 |
apply simp |
|
5207 |
done |
|
5208 |
||
65336 | 5209 |
lemma word_rec_in: "f (word_rec z (\<lambda>_. f) n) = word_rec (f z) (\<lambda>_. f) n" |
37660 | 5210 |
by (induct n) (simp_all add: word_rec_0 word_rec_Suc) |
5211 |
||
67399 | 5212 |
lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \<circ> (+) 1) n" |
37660 | 5213 |
by (induct n) (simp_all add: word_rec_0 word_rec_Suc) |
5214 |
||
65268 | 5215 |
lemma word_rec_twice: |
67399 | 5216 |
"m \<le> n \<Longrightarrow> word_rec z f n = word_rec (word_rec z f (n - m)) (f \<circ> (+) (n - m)) m" |
65336 | 5217 |
apply (erule rev_mp) |
5218 |
apply (rule_tac x=z in spec) |
|
5219 |
apply (rule_tac x=f in spec) |
|
5220 |
apply (induct n) |
|
5221 |
apply (simp add: word_rec_0) |
|
5222 |
apply clarsimp |
|
5223 |
apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) |
|
5224 |
apply simp |
|
5225 |
apply (case_tac "1 + (n - m) = 0") |
|
5226 |
apply (simp add: word_rec_0) |
|
5227 |
apply (rule_tac f = "word_rec a b" for a b in arg_cong) |
|
5228 |
apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) |
|
5229 |
apply simp |
|
5230 |
apply (simp (no_asm_use)) |
|
5231 |
apply (simp add: word_rec_Suc word_rec_in2) |
|
5232 |
apply (erule impE) |
|
5233 |
apply uint_arith |
|
67399 | 5234 |
apply (drule_tac x="x \<circ> (+) 1" in spec) |
65336 | 5235 |
apply (drule_tac x="x 0 xa" in spec) |
37660 | 5236 |
apply simp |
65336 | 5237 |
apply (rule_tac t="\<lambda>a. x (1 + (n - m + a))" and s="\<lambda>a. x (1 + (n - m) + a)" in subst) |
5238 |
apply (clarsimp simp add: fun_eq_iff) |
|
5239 |
apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) |
|
5240 |
apply simp |
|
5241 |
apply (rule refl) |
|
5242 |
apply (rule refl) |
|
5243 |
done |
|
37660 | 5244 |
|
5245 |
lemma word_rec_id: "word_rec z (\<lambda>_. id) n = z" |
|
5246 |
by (induct n) (auto simp add: word_rec_0 word_rec_Suc) |
|
5247 |
||
5248 |
lemma word_rec_id_eq: "\<forall>m < n. f m = id \<Longrightarrow> word_rec z f n = z" |
|
65336 | 5249 |
apply (erule rev_mp) |
5250 |
apply (induct n) |
|
5251 |
apply (auto simp add: word_rec_0 word_rec_Suc) |
|
5252 |
apply (drule spec, erule mp) |
|
5253 |
apply uint_arith |
|
5254 |
apply (drule_tac x=n in spec, erule impE) |
|
5255 |
apply uint_arith |
|
5256 |
apply simp |
|
5257 |
done |
|
37660 | 5258 |
|
65268 | 5259 |
lemma word_rec_max: |
58410
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents:
58061
diff
changeset
|
5260 |
"\<forall>m\<ge>n. m \<noteq> - 1 \<longrightarrow> f m = id \<Longrightarrow> word_rec z f (- 1) = word_rec z f n" |
65336 | 5261 |
apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) |
5262 |
apply simp |
|
5263 |
apply simp |
|
5264 |
apply (rule word_rec_id_eq) |
|
5265 |
apply clarsimp |
|
5266 |
apply (drule spec, rule mp, erule mp) |
|
5267 |
apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) |
|
5268 |
prefer 2 |
|
5269 |
apply assumption |
|
5270 |
apply simp |
|
5271 |
apply (erule contrapos_pn) |
|
5272 |
apply simp |
|
5273 |
apply (drule arg_cong[where f="\<lambda>x. x - n"]) |
|
5274 |
apply simp |
|
5275 |
done |
|
5276 |
||
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5277 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5278 |
subsection \<open>More\<close> |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5279 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5280 |
lemma test_bit_1' [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
|
5281 |
"(1 :: 'a :: len word) !! n \<longleftrightarrow> 0 < LENGTH('a) \<and> n = 0" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
5282 |
by simp |
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5283 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5284 |
lemma mask_0 [simp]: |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5285 |
"mask 0 = 0" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5286 |
by (simp add: Word.mask_def) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5287 |
|
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
5288 |
lemma shiftl0: |
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
|
5289 |
"x << 0 = (x :: 'a :: len word)" |
71957
3e162c63371a
build bit operations on word on library theory on bit operations
haftmann
parents:
71955
diff
changeset
|
5290 |
by (fact shiftl_x_0) |
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5291 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5292 |
lemma mask_1: "mask 1 = 1" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5293 |
by (simp add: mask_def) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5294 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5295 |
lemma mask_Suc_0: "mask (Suc 0) = 1" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5296 |
by (simp add: mask_def) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5297 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5298 |
lemma mask_numeral: "mask (numeral n) = 2 * mask (pred_numeral n) + 1" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5299 |
by (simp add: mask_def neg_numeral_class.sub_def numeral_eq_Suc numeral_pow) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5300 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5301 |
lemma bin_last_bintrunc: "bin_last (bintrunc l n) = (l > 0 \<and> bin_last n)" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5302 |
by (cases l) simp_all |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5303 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5304 |
lemma word_and_1: |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5305 |
"n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5306 |
by transfer (rule bin_rl_eqI, simp_all add: bin_rest_trunc bin_last_bintrunc) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5307 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5308 |
lemma bintrunc_shiftl: |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5309 |
"bintrunc n (m << i) = bintrunc (n - i) m << i" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5310 |
proof (induction i arbitrary: n) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5311 |
case 0 |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5312 |
show ?case |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5313 |
by simp |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5314 |
next |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5315 |
case (Suc i) |
71986 | 5316 |
then show ?case by (cases n) (simp_all add: take_bit_Suc) |
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5317 |
qed |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5318 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5319 |
lemma shiftl_transfer [transfer_rule]: |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5320 |
includes lifting_syntax |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5321 |
shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5322 |
by (auto intro!: rel_funI word_eqI simp add: word.pcr_cr_eq cr_word_def word_size nth_shiftl) |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5323 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5324 |
lemma uint_shiftl: |
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5325 |
"uint (n << i) = bintrunc (size n) (uint n << i)" |
71990 | 5326 |
apply (simp add: word_size shiftl_eq_push_bit shiftl_word_eq) |
5327 |
apply transfer |
|
5328 |
apply (simp add: push_bit_take_bit) |
|
5329 |
done |
|
70183
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5330 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5331 |
|
3ea80c950023
incorporated various material from the AFP into the distribution
haftmann
parents:
70175
diff
changeset
|
5332 |
subsection \<open>Misc\<close> |
37660 | 5333 |
|
5334 |
declare bin_to_bl_def [simp] |
|
5335 |
||
69605 | 5336 |
ML_file \<open>Tools/word_lib.ML\<close> |
5337 |
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
|
5338 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
5339 |
hide_const (open) Word |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46962
diff
changeset
|
5340 |
|
41060
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
boehmes
parents:
40827
diff
changeset
|
5341 |
end |