| author | blanchet |
| Tue, 05 Oct 2010 11:45:10 +0200 | |
| changeset 39953 | aa54f347e5e2 |
| parent 35416 | d8d7d1b785af |
| child 44763 | b50d5d694838 |
| permissions | -rw-r--r-- |
| 15647 | 1 |
(* AUTOMATICALLY GENERATED, DO NOT EDIT! *) |
2 |
||
| 17566 | 3 |
theory HOL4Vec imports HOL4Base begin |
| 14516 | 4 |
|
5 |
;setup_theory res_quan |
|
6 |
||
| 17644 | 7 |
lemma RES_FORALL_CONJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. |
8 |
RES_FORALL P (%i::'a::type. Q i & R i) = |
|
9 |
(RES_FORALL P Q & RES_FORALL P R)" |
|
| 14516 | 10 |
by (import res_quan RES_FORALL_CONJ_DIST) |
11 |
||
| 17644 | 12 |
lemma RES_FORALL_DISJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. |
13 |
RES_FORALL (%j::'a::type. P j | Q j) R = |
|
14 |
(RES_FORALL P R & RES_FORALL Q R)" |
|
| 14516 | 15 |
by (import res_quan RES_FORALL_DISJ_DIST) |
16 |
||
| 17644 | 17 |
lemma RES_FORALL_UNIQUE: "ALL (x::'a::type => bool) xa::'a::type. RES_FORALL (op = xa) x = x xa" |
| 14516 | 18 |
by (import res_quan RES_FORALL_UNIQUE) |
19 |
||
| 17644 | 20 |
lemma RES_FORALL_FORALL: "ALL (P::'a::type => bool) (R::'a::type => 'b::type => bool) x::'b::type. |
21 |
(ALL x::'b::type. RES_FORALL P (%i::'a::type. R i x)) = |
|
22 |
RES_FORALL P (%i::'a::type. All (R i))" |
|
| 14516 | 23 |
by (import res_quan RES_FORALL_FORALL) |
24 |
||
| 17644 | 25 |
lemma RES_FORALL_REORDER: "ALL (P::'a::type => bool) (Q::'b::type => bool) |
26 |
R::'a::type => 'b::type => bool. |
|
27 |
RES_FORALL P (%i::'a::type. RES_FORALL Q (R i)) = |
|
28 |
RES_FORALL Q (%j::'b::type. RES_FORALL P (%i::'a::type. R i j))" |
|
| 14516 | 29 |
by (import res_quan RES_FORALL_REORDER) |
30 |
||
31 |
lemma RES_FORALL_EMPTY: "All (RES_FORALL EMPTY)" |
|
32 |
by (import res_quan RES_FORALL_EMPTY) |
|
33 |
||
| 17644 | 34 |
lemma RES_FORALL_UNIV: "ALL p::'a::type => bool. RES_FORALL pred_set.UNIV p = All p" |
| 14516 | 35 |
by (import res_quan RES_FORALL_UNIV) |
36 |
||
| 17644 | 37 |
lemma RES_FORALL_NULL: "ALL (p::'a::type => bool) m::bool. |
38 |
RES_FORALL p (%x::'a::type. m) = (p = EMPTY | m)" |
|
| 14516 | 39 |
by (import res_quan RES_FORALL_NULL) |
40 |
||
| 17644 | 41 |
lemma RES_EXISTS_DISJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. |
42 |
RES_EXISTS P (%i::'a::type. Q i | R i) = |
|
43 |
(RES_EXISTS P Q | RES_EXISTS P R)" |
|
| 14516 | 44 |
by (import res_quan RES_EXISTS_DISJ_DIST) |
45 |
||
| 17644 | 46 |
lemma RES_DISJ_EXISTS_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. |
47 |
RES_EXISTS (%i::'a::type. P i | Q i) R = |
|
48 |
(RES_EXISTS P R | RES_EXISTS Q R)" |
|
| 14516 | 49 |
by (import res_quan RES_DISJ_EXISTS_DIST) |
50 |
||
| 17644 | 51 |
lemma RES_EXISTS_EQUAL: "ALL (x::'a::type => bool) xa::'a::type. RES_EXISTS (op = xa) x = x xa" |
| 14516 | 52 |
by (import res_quan RES_EXISTS_EQUAL) |
53 |
||
| 17644 | 54 |
lemma RES_EXISTS_REORDER: "ALL (P::'a::type => bool) (Q::'b::type => bool) |
55 |
R::'a::type => 'b::type => bool. |
|
56 |
RES_EXISTS P (%i::'a::type. RES_EXISTS Q (R i)) = |
|
57 |
RES_EXISTS Q (%j::'b::type. RES_EXISTS P (%i::'a::type. R i j))" |
|
| 14516 | 58 |
by (import res_quan RES_EXISTS_REORDER) |
59 |
||
| 17644 | 60 |
lemma RES_EXISTS_EMPTY: "ALL p::'a::type => bool. ~ RES_EXISTS EMPTY p" |
| 14516 | 61 |
by (import res_quan RES_EXISTS_EMPTY) |
62 |
||
| 17644 | 63 |
lemma RES_EXISTS_UNIV: "ALL p::'a::type => bool. RES_EXISTS pred_set.UNIV p = Ex p" |
| 14516 | 64 |
by (import res_quan RES_EXISTS_UNIV) |
65 |
||
| 17644 | 66 |
lemma RES_EXISTS_NULL: "ALL (p::'a::type => bool) m::bool. |
67 |
RES_EXISTS p (%x::'a::type. m) = (p ~= EMPTY & m)" |
|
| 14516 | 68 |
by (import res_quan RES_EXISTS_NULL) |
69 |
||
| 17644 | 70 |
lemma RES_EXISTS_ALT: "ALL (p::'a::type => bool) m::'a::type => bool. |
71 |
RES_EXISTS p m = (IN (RES_SELECT p m) p & m (RES_SELECT p m))" |
|
| 14516 | 72 |
by (import res_quan RES_EXISTS_ALT) |
73 |
||
| 17644 | 74 |
lemma RES_EXISTS_UNIQUE_EMPTY: "ALL p::'a::type => bool. ~ RES_EXISTS_UNIQUE EMPTY p" |
| 14516 | 75 |
by (import res_quan RES_EXISTS_UNIQUE_EMPTY) |
76 |
||
| 17644 | 77 |
lemma RES_EXISTS_UNIQUE_UNIV: "ALL p::'a::type => bool. RES_EXISTS_UNIQUE pred_set.UNIV p = Ex1 p" |
| 14516 | 78 |
by (import res_quan RES_EXISTS_UNIQUE_UNIV) |
79 |
||
| 17644 | 80 |
lemma RES_EXISTS_UNIQUE_NULL: "ALL (p::'a::type => bool) m::bool. |
81 |
RES_EXISTS_UNIQUE p (%x::'a::type. m) = |
|
82 |
((EX x::'a::type. p = INSERT x EMPTY) & m)" |
|
| 14516 | 83 |
by (import res_quan RES_EXISTS_UNIQUE_NULL) |
84 |
||
| 17644 | 85 |
lemma RES_EXISTS_UNIQUE_ALT: "ALL (p::'a::type => bool) m::'a::type => bool. |
| 14516 | 86 |
RES_EXISTS_UNIQUE p m = |
| 17644 | 87 |
RES_EXISTS p |
88 |
(%x::'a::type. m x & RES_FORALL p (%y::'a::type. m y --> y = x))" |
|
| 14516 | 89 |
by (import res_quan RES_EXISTS_UNIQUE_ALT) |
90 |
||
| 17644 | 91 |
lemma RES_SELECT_EMPTY: "ALL p::'a::type => bool. RES_SELECT EMPTY p = (SOME x::'a::type. False)" |
| 14516 | 92 |
by (import res_quan RES_SELECT_EMPTY) |
93 |
||
| 17644 | 94 |
lemma RES_SELECT_UNIV: "ALL p::'a::type => bool. RES_SELECT pred_set.UNIV p = Eps p" |
| 14516 | 95 |
by (import res_quan RES_SELECT_UNIV) |
96 |
||
| 17644 | 97 |
lemma RES_ABSTRACT: "ALL (p::'a::type => bool) (m::'a::type => 'b::type) x::'a::type. |
98 |
IN x p --> RES_ABSTRACT p m x = m x" |
|
| 14516 | 99 |
by (import res_quan RES_ABSTRACT) |
100 |
||
| 17644 | 101 |
lemma RES_ABSTRACT_EQUAL: "ALL (p::'a::type => bool) (m1::'a::type => 'b::type) |
102 |
m2::'a::type => 'b::type. |
|
103 |
(ALL x::'a::type. IN x p --> m1 x = m2 x) --> |
|
104 |
RES_ABSTRACT p m1 = RES_ABSTRACT p m2" |
|
| 14516 | 105 |
by (import res_quan RES_ABSTRACT_EQUAL) |
106 |
||
| 17644 | 107 |
lemma RES_ABSTRACT_IDEMPOT: "ALL (p::'a::type => bool) m::'a::type => 'b::type. |
108 |
RES_ABSTRACT p (RES_ABSTRACT p m) = RES_ABSTRACT p m" |
|
| 14516 | 109 |
by (import res_quan RES_ABSTRACT_IDEMPOT) |
110 |
||
| 17644 | 111 |
lemma RES_ABSTRACT_EQUAL_EQ: "ALL (p::'a::type => bool) (m1::'a::type => 'b::type) |
112 |
m2::'a::type => 'b::type. |
|
113 |
(RES_ABSTRACT p m1 = RES_ABSTRACT p m2) = |
|
114 |
(ALL x::'a::type. IN x p --> m1 x = m2 x)" |
|
| 14516 | 115 |
by (import res_quan RES_ABSTRACT_EQUAL_EQ) |
116 |
||
117 |
;end_setup |
|
118 |
||
119 |
;setup_theory word_base |
|
120 |
||
| 17644 | 121 |
typedef (open) ('a) word = "(Collect::('a::type list recspace => bool) => 'a::type list recspace set)
|
122 |
(%x::'a::type list recspace. |
|
123 |
(All::(('a::type list recspace => bool) => bool) => bool)
|
|
124 |
(%word::'a::type list recspace => bool. |
|
| 14516 | 125 |
(op -->::bool => bool => bool) |
| 17644 | 126 |
((All::('a::type list recspace => bool) => bool)
|
127 |
(%a0::'a::type list recspace. |
|
| 14516 | 128 |
(op -->::bool => bool => bool) |
| 17644 | 129 |
((Ex::('a::type list => bool) => bool)
|
130 |
(%a::'a::type list. |
|
131 |
(op =::'a::type list recspace |
|
132 |
=> 'a::type list recspace => bool) |
|
| 14516 | 133 |
a0 ((CONSTR::nat |
| 17644 | 134 |
=> 'a::type list |
135 |
=> (nat => 'a::type list recspace) => 'a::type list recspace) |
|
| 14516 | 136 |
(0::nat) a |
| 17644 | 137 |
(%n::nat. BOTTOM::'a::type list recspace)))) |
| 14516 | 138 |
(word a0))) |
139 |
(word x)))" |
|
140 |
by (rule typedef_helper,import word_base word_TY_DEF) |
|
141 |
||
142 |
lemmas word_TY_DEF = typedef_hol2hol4 [OF type_definition_word] |
|
143 |
||
144 |
consts |
|
| 17652 | 145 |
mk_word :: "'a list recspace => 'a word" |
146 |
dest_word :: "'a word => 'a list recspace" |
|
| 14516 | 147 |
|
| 17644 | 148 |
specification (dest_word mk_word) word_repfns: "(ALL a::'a::type word. mk_word (dest_word a) = a) & |
149 |
(ALL r::'a::type list recspace. |
|
150 |
(ALL word::'a::type list recspace => bool. |
|
151 |
(ALL a0::'a::type list recspace. |
|
| 17652 | 152 |
(EX a::'a::type list. a0 = CONSTR 0 a (%n::nat. BOTTOM)) --> |
| 14516 | 153 |
word a0) --> |
154 |
word r) = |
|
155 |
(dest_word (mk_word r) = r))" |
|
156 |
by (import word_base word_repfns) |
|
157 |
||
158 |
consts |
|
| 17652 | 159 |
word_base0 :: "'a list => 'a word" |
| 14516 | 160 |
|
161 |
defs |
|
| 17652 | 162 |
word_base0_primdef: "word_base0 == %a::'a::type list. mk_word (CONSTR 0 a (%n::nat. BOTTOM))" |
| 14516 | 163 |
|
| 17652 | 164 |
lemma word_base0_def: "word_base0 = (%a::'a::type list. mk_word (CONSTR 0 a (%n::nat. BOTTOM)))" |
| 14516 | 165 |
by (import word_base word_base0_def) |
166 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
167 |
definition WORD :: "'a list => 'a word" where |
| 14516 | 168 |
"WORD == word_base0" |
169 |
||
170 |
lemma WORD: "WORD = word_base0" |
|
171 |
by (import word_base WORD) |
|
172 |
||
173 |
consts |
|
| 17652 | 174 |
word_case :: "('a list => 'b) => 'a word => 'b"
|
| 14516 | 175 |
|
| 17644 | 176 |
specification (word_case_primdef: word_case) word_case_def: "ALL (f::'a::type list => 'b::type) a::'a::type list. |
177 |
word_case f (WORD a) = f a" |
|
| 14516 | 178 |
by (import word_base word_case_def) |
179 |
||
180 |
consts |
|
| 17652 | 181 |
word_size :: "('a => nat) => 'a word => nat"
|
| 14516 | 182 |
|
| 17644 | 183 |
specification (word_size_primdef: word_size) word_size_def: "ALL (f::'a::type => nat) a::'a::type list. |
| 17652 | 184 |
word_size f (WORD a) = 1 + list_size f a" |
| 14516 | 185 |
by (import word_base word_size_def) |
186 |
||
| 17644 | 187 |
lemma word_11: "ALL (a::'a::type list) a'::'a::type list. (WORD a = WORD a') = (a = a')" |
| 14516 | 188 |
by (import word_base word_11) |
189 |
||
| 17644 | 190 |
lemma word_case_cong: "ALL (M::'a::type word) (M'::'a::type word) f::'a::type list => 'b::type. |
191 |
M = M' & |
|
192 |
(ALL a::'a::type list. |
|
193 |
M' = WORD a --> f a = (f'::'a::type list => 'b::type) a) --> |
|
| 14516 | 194 |
word_case f M = word_case f' M'" |
195 |
by (import word_base word_case_cong) |
|
196 |
||
| 17644 | 197 |
lemma word_nchotomy: "ALL x::'a::type word. EX l::'a::type list. x = WORD l" |
| 14516 | 198 |
by (import word_base word_nchotomy) |
199 |
||
| 17644 | 200 |
lemma word_Axiom: "ALL f::'a::type list => 'b::type. |
201 |
EX fn::'a::type word => 'b::type. ALL a::'a::type list. fn (WORD a) = f a" |
|
| 14516 | 202 |
by (import word_base word_Axiom) |
203 |
||
| 17644 | 204 |
lemma word_induction: "ALL P::'a::type word => bool. (ALL a::'a::type list. P (WORD a)) --> All P" |
| 14516 | 205 |
by (import word_base word_induction) |
206 |
||
| 17644 | 207 |
lemma word_Ax: "ALL f::'a::type list => 'b::type. |
208 |
EX fn::'a::type word => 'b::type. ALL a::'a::type list. fn (WORD a) = f a" |
|
| 14516 | 209 |
by (import word_base word_Ax) |
210 |
||
| 17644 | 211 |
lemma WORD_11: "ALL (x::'a::type list) xa::'a::type list. (WORD x = WORD xa) = (x = xa)" |
| 14516 | 212 |
by (import word_base WORD_11) |
213 |
||
| 17644 | 214 |
lemma word_induct: "ALL x::'a::type word => bool. (ALL l::'a::type list. x (WORD l)) --> All x" |
| 14516 | 215 |
by (import word_base word_induct) |
216 |
||
| 17644 | 217 |
lemma word_cases: "ALL x::'a::type word. EX l::'a::type list. x = WORD l" |
| 14516 | 218 |
by (import word_base word_cases) |
219 |
||
220 |
consts |
|
| 17652 | 221 |
WORDLEN :: "'a word => nat" |
| 14516 | 222 |
|
| 17644 | 223 |
specification (WORDLEN) WORDLEN_DEF: "ALL l::'a::type list. WORDLEN (WORD l) = length l" |
| 14516 | 224 |
by (import word_base WORDLEN_DEF) |
225 |
||
226 |
consts |
|
| 17652 | 227 |
PWORDLEN :: "nat => 'a word => bool" |
| 14516 | 228 |
|
229 |
defs |
|
| 17644 | 230 |
PWORDLEN_primdef: "PWORDLEN == %n::nat. GSPEC (%w::'a::type word. (w, WORDLEN w = n))" |
| 14516 | 231 |
|
| 17644 | 232 |
lemma PWORDLEN_def: "ALL n::nat. PWORDLEN n = GSPEC (%w::'a::type word. (w, WORDLEN w = n))" |
| 14516 | 233 |
by (import word_base PWORDLEN_def) |
234 |
||
| 17644 | 235 |
lemma IN_PWORDLEN: "ALL (n::nat) l::'a::type list. IN (WORD l) (PWORDLEN n) = (length l = n)" |
| 14516 | 236 |
by (import word_base IN_PWORDLEN) |
237 |
||
| 17644 | 238 |
lemma PWORDLEN: "ALL (n::nat) w::'a::type word. IN w (PWORDLEN n) = (WORDLEN w = n)" |
| 14516 | 239 |
by (import word_base PWORDLEN) |
240 |
||
| 17652 | 241 |
lemma PWORDLEN0: "ALL w::'a::type word. IN w (PWORDLEN 0) --> w = WORD []" |
| 14516 | 242 |
by (import word_base PWORDLEN0) |
243 |
||
| 17652 | 244 |
lemma PWORDLEN1: "ALL x::'a::type. IN (WORD [x]) (PWORDLEN 1)" |
| 14516 | 245 |
by (import word_base PWORDLEN1) |
246 |
||
247 |
consts |
|
| 17652 | 248 |
WSEG :: "nat => nat => 'a word => 'a word" |
| 14516 | 249 |
|
| 17644 | 250 |
specification (WSEG) WSEG_DEF: "ALL (m::nat) (k::nat) l::'a::type list. |
251 |
WSEG m k (WORD l) = WORD (LASTN m (BUTLASTN k l))" |
|
| 14516 | 252 |
by (import word_base WSEG_DEF) |
253 |
||
| 17652 | 254 |
lemma WSEG0: "ALL (k::nat) w::'a::type word. WSEG 0 k w = WORD []" |
| 14516 | 255 |
by (import word_base WSEG0) |
256 |
||
| 17644 | 257 |
lemma WSEG_PWORDLEN: "ALL n::nat. |
| 14516 | 258 |
RES_FORALL (PWORDLEN n) |
| 17644 | 259 |
(%w::'a::type word. |
260 |
ALL (m::nat) k::nat. m + k <= n --> IN (WSEG m k w) (PWORDLEN m))" |
|
| 14516 | 261 |
by (import word_base WSEG_PWORDLEN) |
262 |
||
| 17644 | 263 |
lemma WSEG_WORDLEN: "ALL x::nat. |
| 14516 | 264 |
RES_FORALL (PWORDLEN x) |
| 17644 | 265 |
(%xa::'a::type word. |
266 |
ALL (xb::nat) xc::nat. |
|
267 |
xb + xc <= x --> WORDLEN (WSEG xb xc xa) = xb)" |
|
| 14516 | 268 |
by (import word_base WSEG_WORDLEN) |
269 |
||
| 17652 | 270 |
lemma WSEG_WORD_LENGTH: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::'a::type word. WSEG n 0 w = w)" |
| 14516 | 271 |
by (import word_base WSEG_WORD_LENGTH) |
272 |
||
273 |
consts |
|
| 17652 | 274 |
bit :: "nat => 'a word => 'a" |
| 14516 | 275 |
|
| 17644 | 276 |
specification (bit) BIT_DEF: "ALL (k::nat) l::'a::type list. bit k (WORD l) = ELL k l" |
| 14516 | 277 |
by (import word_base BIT_DEF) |
278 |
||
| 17652 | 279 |
lemma BIT0: "ALL x::'a::type. bit 0 (WORD [x]) = x" |
| 14516 | 280 |
by (import word_base BIT0) |
281 |
||
| 14847 | 282 |
lemma WSEG_BIT: "(All::(nat => bool) => bool) |
283 |
(%n::nat. |
|
| 17644 | 284 |
(RES_FORALL::('a::type word => bool)
|
285 |
=> ('a::type word => bool) => bool)
|
|
286 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
287 |
(%w::'a::type word. |
|
| 14847 | 288 |
(All::(nat => bool) => bool) |
289 |
(%k::nat. |
|
290 |
(op -->::bool => bool => bool) |
|
291 |
((op <::nat => nat => bool) k n) |
|
| 17644 | 292 |
((op =::'a::type word => 'a::type word => bool) |
293 |
((WSEG::nat => nat => 'a::type word => 'a::type word) |
|
294 |
(1::nat) k w) |
|
295 |
((WORD::'a::type list => 'a::type word) |
|
296 |
((op #::'a::type => 'a::type list => 'a::type list) |
|
297 |
((bit::nat => 'a::type word => 'a::type) k w) |
|
298 |
([]::'a::type list)))))))" |
|
| 14516 | 299 |
by (import word_base WSEG_BIT) |
300 |
||
| 17644 | 301 |
lemma BIT_WSEG: "ALL n::nat. |
| 14516 | 302 |
RES_FORALL (PWORDLEN n) |
| 17644 | 303 |
(%w::'a::type word. |
304 |
ALL (m::nat) (k::nat) j::nat. |
|
305 |
m + k <= n --> j < m --> bit j (WSEG m k w) = bit (j + k) w)" |
|
| 14516 | 306 |
by (import word_base BIT_WSEG) |
307 |
||
308 |
consts |
|
| 17652 | 309 |
MSB :: "'a word => 'a" |
| 14516 | 310 |
|
| 17644 | 311 |
specification (MSB) MSB_DEF: "ALL l::'a::type list. MSB (WORD l) = hd l" |
| 14516 | 312 |
by (import word_base MSB_DEF) |
313 |
||
| 17644 | 314 |
lemma MSB: "ALL n::nat. |
315 |
RES_FORALL (PWORDLEN n) |
|
| 17652 | 316 |
(%w::'a::type word. 0 < n --> MSB w = bit (PRE n) w)" |
| 14516 | 317 |
by (import word_base MSB) |
318 |
||
319 |
consts |
|
| 17652 | 320 |
LSB :: "'a word => 'a" |
| 14516 | 321 |
|
| 17644 | 322 |
specification (LSB) LSB_DEF: "ALL l::'a::type list. LSB (WORD l) = last l" |
| 14516 | 323 |
by (import word_base LSB_DEF) |
324 |
||
| 17644 | 325 |
lemma LSB: "ALL n::nat. |
| 17652 | 326 |
RES_FORALL (PWORDLEN n) (%w::'a::type word. 0 < n --> LSB w = bit 0 w)" |
| 14516 | 327 |
by (import word_base LSB) |
328 |
||
329 |
consts |
|
| 17652 | 330 |
WSPLIT :: "nat => 'a word => 'a word * 'a word" |
| 14516 | 331 |
|
| 17644 | 332 |
specification (WSPLIT) WSPLIT_DEF: "ALL (m::nat) l::'a::type list. |
333 |
WSPLIT m (WORD l) = (WORD (BUTLASTN m l), WORD (LASTN m l))" |
|
| 14516 | 334 |
by (import word_base WSPLIT_DEF) |
335 |
||
336 |
consts |
|
| 17652 | 337 |
WCAT :: "'a word * 'a word => 'a word" |
| 14516 | 338 |
|
| 17644 | 339 |
specification (WCAT) WCAT_DEF: "ALL (l1::'a::type list) l2::'a::type list. |
340 |
WCAT (WORD l1, WORD l2) = WORD (l1 @ l2)" |
|
| 14516 | 341 |
by (import word_base WCAT_DEF) |
342 |
||
| 14847 | 343 |
lemma WORD_PARTITION: "(op &::bool => bool => bool) |
344 |
((All::(nat => bool) => bool) |
|
345 |
(%n::nat. |
|
| 17644 | 346 |
(RES_FORALL::('a::type word => bool)
|
347 |
=> ('a::type word => bool) => bool)
|
|
348 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
349 |
(%w::'a::type word. |
|
| 14847 | 350 |
(All::(nat => bool) => bool) |
351 |
(%m::nat. |
|
352 |
(op -->::bool => bool => bool) |
|
353 |
((op <=::nat => nat => bool) m n) |
|
| 17644 | 354 |
((op =::'a::type word => 'a::type word => bool) |
355 |
((WCAT::'a::type word * 'a::type word => 'a::type word) |
|
356 |
((WSPLIT::nat |
|
357 |
=> 'a::type word |
|
358 |
=> 'a::type word * 'a::type word) |
|
359 |
m w)) |
|
| 14847 | 360 |
w))))) |
361 |
((All::(nat => bool) => bool) |
|
362 |
(%n::nat. |
|
363 |
(All::(nat => bool) => bool) |
|
364 |
(%m::nat. |
|
| 17644 | 365 |
(RES_FORALL::('a::type word => bool)
|
366 |
=> ('a::type word => bool) => bool)
|
|
367 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
368 |
(%w1::'a::type word. |
|
369 |
(RES_FORALL::('a::type word => bool)
|
|
370 |
=> ('a::type word => bool) => bool)
|
|
371 |
((PWORDLEN::nat => 'a::type word => bool) m) |
|
372 |
(%w2::'a::type word. |
|
373 |
(op =::'a::type word * 'a::type word |
|
374 |
=> 'a::type word * 'a::type word => bool) |
|
375 |
((WSPLIT::nat |
|
376 |
=> 'a::type word |
|
377 |
=> 'a::type word * 'a::type word) |
|
378 |
m ((WCAT::'a::type word * 'a::type word |
|
379 |
=> 'a::type word) |
|
380 |
((Pair::'a::type word |
|
381 |
=> 'a::type word |
|
382 |
=> 'a::type word * 'a::type word) |
|
383 |
w1 w2))) |
|
384 |
((Pair::'a::type word |
|
385 |
=> 'a::type word |
|
386 |
=> 'a::type word * 'a::type word) |
|
387 |
w1 w2))))))" |
|
| 14516 | 388 |
by (import word_base WORD_PARTITION) |
389 |
||
| 17644 | 390 |
lemma WCAT_ASSOC: "ALL (w1::'a::type word) (w2::'a::type word) w3::'a::type word. |
391 |
WCAT (w1, WCAT (w2, w3)) = WCAT (WCAT (w1, w2), w3)" |
|
| 14516 | 392 |
by (import word_base WCAT_ASSOC) |
393 |
||
| 17644 | 394 |
lemma WCAT0: "ALL w::'a::type word. WCAT (WORD [], w) = w & WCAT (w, WORD []) = w" |
| 14516 | 395 |
by (import word_base WCAT0) |
396 |
||
| 17644 | 397 |
lemma WCAT_11: "ALL (m::nat) n::nat. |
| 14516 | 398 |
RES_FORALL (PWORDLEN m) |
| 17644 | 399 |
(%wm1::'a::type word. |
400 |
RES_FORALL (PWORDLEN m) |
|
401 |
(%wm2::'a::type word. |
|
402 |
RES_FORALL (PWORDLEN n) |
|
403 |
(%wn1::'a::type word. |
|
404 |
RES_FORALL (PWORDLEN n) |
|
405 |
(%wn2::'a::type word. |
|
406 |
(WCAT (wm1, wn1) = WCAT (wm2, wn2)) = |
|
407 |
(wm1 = wm2 & wn1 = wn2)))))" |
|
| 14516 | 408 |
by (import word_base WCAT_11) |
409 |
||
| 14847 | 410 |
lemma WSPLIT_PWORDLEN: "(All::(nat => bool) => bool) |
411 |
(%n::nat. |
|
| 17644 | 412 |
(RES_FORALL::('a::type word => bool)
|
413 |
=> ('a::type word => bool) => bool)
|
|
414 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
415 |
(%w::'a::type word. |
|
| 14847 | 416 |
(All::(nat => bool) => bool) |
417 |
(%m::nat. |
|
418 |
(op -->::bool => bool => bool) |
|
419 |
((op <=::nat => nat => bool) m n) |
|
420 |
((op &::bool => bool => bool) |
|
| 17644 | 421 |
((IN::'a::type word => ('a::type word => bool) => bool)
|
422 |
((fst::'a::type word * 'a::type word => 'a::type word) |
|
423 |
((WSPLIT::nat |
|
424 |
=> 'a::type word |
|
425 |
=> 'a::type word * 'a::type word) |
|
426 |
m w)) |
|
427 |
((PWORDLEN::nat => 'a::type word => bool) |
|
| 14847 | 428 |
((op -::nat => nat => nat) n m))) |
| 17644 | 429 |
((IN::'a::type word => ('a::type word => bool) => bool)
|
430 |
((snd::'a::type word * 'a::type word => 'a::type word) |
|
431 |
((WSPLIT::nat |
|
432 |
=> 'a::type word |
|
433 |
=> 'a::type word * 'a::type word) |
|
434 |
m w)) |
|
435 |
((PWORDLEN::nat => 'a::type word => bool) m))))))" |
|
| 14516 | 436 |
by (import word_base WSPLIT_PWORDLEN) |
437 |
||
| 17644 | 438 |
lemma WCAT_PWORDLEN: "ALL n1::nat. |
| 14516 | 439 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 440 |
(%w1::'a::type word. |
441 |
ALL n2::nat. |
|
442 |
RES_FORALL (PWORDLEN n2) |
|
443 |
(%w2::'a::type word. IN (WCAT (w1, w2)) (PWORDLEN (n1 + n2))))" |
|
| 14516 | 444 |
by (import word_base WCAT_PWORDLEN) |
445 |
||
| 17644 | 446 |
lemma WORDLEN_SUC_WCAT: "ALL (n::nat) w::'a::type word. |
| 14516 | 447 |
IN w (PWORDLEN (Suc n)) --> |
| 17652 | 448 |
RES_EXISTS (PWORDLEN 1) |
| 17644 | 449 |
(%b::'a::type word. |
450 |
RES_EXISTS (PWORDLEN n) (%w'::'a::type word. w = WCAT (b, w')))" |
|
| 14516 | 451 |
by (import word_base WORDLEN_SUC_WCAT) |
452 |
||
| 17644 | 453 |
lemma WSEG_WSEG: "ALL n::nat. |
| 14516 | 454 |
RES_FORALL (PWORDLEN n) |
| 17644 | 455 |
(%w::'a::type word. |
456 |
ALL (m1::nat) (k1::nat) (m2::nat) k2::nat. |
|
457 |
m1 + k1 <= n & m2 + k2 <= m1 --> |
|
458 |
WSEG m2 k2 (WSEG m1 k1 w) = WSEG m2 (k1 + k2) w)" |
|
| 14516 | 459 |
by (import word_base WSEG_WSEG) |
460 |
||
| 14847 | 461 |
lemma WSPLIT_WSEG: "(All::(nat => bool) => bool) |
462 |
(%n::nat. |
|
| 17644 | 463 |
(RES_FORALL::('a::type word => bool)
|
464 |
=> ('a::type word => bool) => bool)
|
|
465 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
466 |
(%w::'a::type word. |
|
| 14847 | 467 |
(All::(nat => bool) => bool) |
468 |
(%k::nat. |
|
469 |
(op -->::bool => bool => bool) |
|
470 |
((op <=::nat => nat => bool) k n) |
|
| 17644 | 471 |
((op =::'a::type word * 'a::type word |
472 |
=> 'a::type word * 'a::type word => bool) |
|
473 |
((WSPLIT::nat |
|
474 |
=> 'a::type word |
|
475 |
=> 'a::type word * 'a::type word) |
|
476 |
k w) |
|
477 |
((Pair::'a::type word |
|
478 |
=> 'a::type word => 'a::type word * 'a::type word) |
|
479 |
((WSEG::nat => nat => 'a::type word => 'a::type word) |
|
| 14847 | 480 |
((op -::nat => nat => nat) n k) k w) |
| 17644 | 481 |
((WSEG::nat => nat => 'a::type word => 'a::type word) k |
482 |
(0::nat) w))))))" |
|
| 14516 | 483 |
by (import word_base WSPLIT_WSEG) |
484 |
||
| 14847 | 485 |
lemma WSPLIT_WSEG1: "(All::(nat => bool) => bool) |
486 |
(%n::nat. |
|
| 17644 | 487 |
(RES_FORALL::('a::type word => bool)
|
488 |
=> ('a::type word => bool) => bool)
|
|
489 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
490 |
(%w::'a::type word. |
|
| 14847 | 491 |
(All::(nat => bool) => bool) |
492 |
(%k::nat. |
|
493 |
(op -->::bool => bool => bool) |
|
494 |
((op <=::nat => nat => bool) k n) |
|
| 17644 | 495 |
((op =::'a::type word => 'a::type word => bool) |
496 |
((fst::'a::type word * 'a::type word => 'a::type word) |
|
497 |
((WSPLIT::nat |
|
498 |
=> 'a::type word |
|
499 |
=> 'a::type word * 'a::type word) |
|
500 |
k w)) |
|
501 |
((WSEG::nat => nat => 'a::type word => 'a::type word) |
|
| 14847 | 502 |
((op -::nat => nat => nat) n k) k w)))))" |
| 14516 | 503 |
by (import word_base WSPLIT_WSEG1) |
504 |
||
| 14847 | 505 |
lemma WSPLIT_WSEG2: "(All::(nat => bool) => bool) |
506 |
(%n::nat. |
|
| 17644 | 507 |
(RES_FORALL::('a::type word => bool)
|
508 |
=> ('a::type word => bool) => bool)
|
|
509 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
510 |
(%w::'a::type word. |
|
| 14847 | 511 |
(All::(nat => bool) => bool) |
512 |
(%k::nat. |
|
513 |
(op -->::bool => bool => bool) |
|
514 |
((op <=::nat => nat => bool) k n) |
|
| 17644 | 515 |
((op =::'a::type word => 'a::type word => bool) |
516 |
((snd::'a::type word * 'a::type word => 'a::type word) |
|
517 |
((WSPLIT::nat |
|
518 |
=> 'a::type word |
|
519 |
=> 'a::type word * 'a::type word) |
|
520 |
k w)) |
|
521 |
((WSEG::nat => nat => 'a::type word => 'a::type word) k |
|
522 |
(0::nat) w)))))" |
|
| 14516 | 523 |
by (import word_base WSPLIT_WSEG2) |
524 |
||
| 17644 | 525 |
lemma WCAT_WSEG_WSEG: "ALL n::nat. |
| 14516 | 526 |
RES_FORALL (PWORDLEN n) |
| 17644 | 527 |
(%w::'a::type word. |
528 |
ALL (m1::nat) (m2::nat) k::nat. |
|
529 |
m1 + (m2 + k) <= n --> |
|
530 |
WCAT (WSEG m2 (m1 + k) w, WSEG m1 k w) = WSEG (m1 + m2) k w)" |
|
| 14516 | 531 |
by (import word_base WCAT_WSEG_WSEG) |
532 |
||
| 17644 | 533 |
lemma WORD_SPLIT: "ALL (x::nat) xa::nat. |
534 |
RES_FORALL (PWORDLEN (x + xa)) |
|
| 17652 | 535 |
(%w::'a::type word. w = WCAT (WSEG x xa w, WSEG xa 0 w))" |
| 14516 | 536 |
by (import word_base WORD_SPLIT) |
537 |
||
| 17644 | 538 |
lemma WORDLEN_SUC_WCAT_WSEG_WSEG: "RES_FORALL (PWORDLEN (Suc (n::nat))) |
| 17652 | 539 |
(%w::'a::type word. w = WCAT (WSEG 1 n w, WSEG n 0 w))" |
| 14516 | 540 |
by (import word_base WORDLEN_SUC_WCAT_WSEG_WSEG) |
541 |
||
| 17644 | 542 |
lemma WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT: "RES_FORALL (PWORDLEN (Suc (n::nat))) |
| 17652 | 543 |
(%w::'a::type word. w = WCAT (WSEG n 1 w, WSEG 1 0 w))" |
| 14516 | 544 |
by (import word_base WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT) |
545 |
||
| 17644 | 546 |
lemma WORDLEN_SUC_WCAT_BIT_WSEG: "ALL n::nat. |
547 |
RES_FORALL (PWORDLEN (Suc n)) |
|
| 17652 | 548 |
(%w::'a::type word. w = WCAT (WORD [bit n w], WSEG n 0 w))" |
| 14516 | 549 |
by (import word_base WORDLEN_SUC_WCAT_BIT_WSEG) |
550 |
||
| 17644 | 551 |
lemma WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT: "ALL n::nat. |
552 |
RES_FORALL (PWORDLEN (Suc n)) |
|
| 17652 | 553 |
(%w::'a::type word. w = WCAT (WSEG n 1 w, WORD [bit 0 w]))" |
| 14516 | 554 |
by (import word_base WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT) |
555 |
||
| 17644 | 556 |
lemma WSEG_WCAT1: "ALL (n1::nat) n2::nat. |
| 14516 | 557 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 558 |
(%w1::'a::type word. |
559 |
RES_FORALL (PWORDLEN n2) |
|
560 |
(%w2::'a::type word. WSEG n1 n2 (WCAT (w1, w2)) = w1))" |
|
| 14516 | 561 |
by (import word_base WSEG_WCAT1) |
562 |
||
| 17644 | 563 |
lemma WSEG_WCAT2: "ALL (n1::nat) n2::nat. |
| 14516 | 564 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 565 |
(%w1::'a::type word. |
566 |
RES_FORALL (PWORDLEN n2) |
|
| 17652 | 567 |
(%w2::'a::type word. WSEG n2 0 (WCAT (w1, w2)) = w2))" |
| 14516 | 568 |
by (import word_base WSEG_WCAT2) |
569 |
||
| 17644 | 570 |
lemma WSEG_SUC: "ALL n::nat. |
| 14516 | 571 |
RES_FORALL (PWORDLEN n) |
| 17644 | 572 |
(%w::'a::type word. |
573 |
ALL (k::nat) m1::nat. |
|
574 |
k + Suc m1 < n --> |
|
| 17652 | 575 |
WSEG (Suc m1) k w = WCAT (WSEG 1 (k + m1) w, WSEG m1 k w))" |
| 14516 | 576 |
by (import word_base WSEG_SUC) |
577 |
||
| 17644 | 578 |
lemma WORD_CONS_WCAT: "ALL (x::'a::type) l::'a::type list. WORD (x # l) = WCAT (WORD [x], WORD l)" |
| 14516 | 579 |
by (import word_base WORD_CONS_WCAT) |
580 |
||
| 17644 | 581 |
lemma WORD_SNOC_WCAT: "ALL (l::'a::type list) x::'a::type. |
582 |
WORD (SNOC x l) = WCAT (WORD l, WORD [x])" |
|
| 14516 | 583 |
by (import word_base WORD_SNOC_WCAT) |
584 |
||
| 17644 | 585 |
lemma BIT_WCAT_FST: "ALL (n1::nat) n2::nat. |
| 14516 | 586 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 587 |
(%w1::'a::type word. |
588 |
RES_FORALL (PWORDLEN n2) |
|
589 |
(%w2::'a::type word. |
|
590 |
ALL k::nat. |
|
591 |
n2 <= k & k < n1 + n2 --> |
|
592 |
bit k (WCAT (w1, w2)) = bit (k - n2) w1))" |
|
| 14516 | 593 |
by (import word_base BIT_WCAT_FST) |
594 |
||
| 14847 | 595 |
lemma BIT_WCAT_SND: "(All::(nat => bool) => bool) |
596 |
(%n1::nat. |
|
597 |
(All::(nat => bool) => bool) |
|
598 |
(%n2::nat. |
|
| 17644 | 599 |
(RES_FORALL::('a::type word => bool)
|
600 |
=> ('a::type word => bool) => bool)
|
|
601 |
((PWORDLEN::nat => 'a::type word => bool) n1) |
|
602 |
(%w1::'a::type word. |
|
603 |
(RES_FORALL::('a::type word => bool)
|
|
604 |
=> ('a::type word => bool) => bool)
|
|
605 |
((PWORDLEN::nat => 'a::type word => bool) n2) |
|
606 |
(%w2::'a::type word. |
|
| 14847 | 607 |
(All::(nat => bool) => bool) |
608 |
(%k::nat. |
|
609 |
(op -->::bool => bool => bool) |
|
610 |
((op <::nat => nat => bool) k n2) |
|
| 17644 | 611 |
((op =::'a::type => 'a::type => bool) |
612 |
((bit::nat => 'a::type word => 'a::type) k |
|
613 |
((WCAT::'a::type word * 'a::type word |
|
614 |
=> 'a::type word) |
|
615 |
((Pair::'a::type word |
|
616 |
=> 'a::type word => 'a::type word * 'a::type word) |
|
| 14847 | 617 |
w1 w2))) |
| 17644 | 618 |
((bit::nat => 'a::type word => 'a::type) k |
619 |
w2)))))))" |
|
| 14516 | 620 |
by (import word_base BIT_WCAT_SND) |
621 |
||
| 17644 | 622 |
lemma BIT_WCAT1: "ALL n::nat. |
623 |
RES_FORALL (PWORDLEN n) |
|
624 |
(%w::'a::type word. ALL b::'a::type. bit n (WCAT (WORD [b], w)) = b)" |
|
| 14516 | 625 |
by (import word_base BIT_WCAT1) |
626 |
||
| 17644 | 627 |
lemma WSEG_WCAT_WSEG1: "ALL (n1::nat) n2::nat. |
| 14516 | 628 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 629 |
(%w1::'a::type word. |
630 |
RES_FORALL (PWORDLEN n2) |
|
631 |
(%w2::'a::type word. |
|
632 |
ALL (m::nat) k::nat. |
|
633 |
m <= n1 & n2 <= k --> |
|
634 |
WSEG m k (WCAT (w1, w2)) = WSEG m (k - n2) w1))" |
|
| 14516 | 635 |
by (import word_base WSEG_WCAT_WSEG1) |
636 |
||
| 17644 | 637 |
lemma WSEG_WCAT_WSEG2: "ALL (n1::nat) n2::nat. |
| 14516 | 638 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 639 |
(%w1::'a::type word. |
640 |
RES_FORALL (PWORDLEN n2) |
|
641 |
(%w2::'a::type word. |
|
642 |
ALL (m::nat) k::nat. |
|
643 |
m + k <= n2 --> WSEG m k (WCAT (w1, w2)) = WSEG m k w2))" |
|
| 14516 | 644 |
by (import word_base WSEG_WCAT_WSEG2) |
645 |
||
| 17644 | 646 |
lemma WSEG_WCAT_WSEG: "ALL (n1::nat) n2::nat. |
| 14516 | 647 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 648 |
(%w1::'a::type word. |
649 |
RES_FORALL (PWORDLEN n2) |
|
650 |
(%w2::'a::type word. |
|
651 |
ALL (m::nat) k::nat. |
|
652 |
m + k <= n1 + n2 & k < n2 & n2 <= m + k --> |
|
653 |
WSEG m k (WCAT (w1, w2)) = |
|
| 17652 | 654 |
WCAT (WSEG (m + k - n2) 0 w1, WSEG (n2 - k) k w2)))" |
| 14516 | 655 |
by (import word_base WSEG_WCAT_WSEG) |
656 |
||
| 14847 | 657 |
lemma BIT_EQ_IMP_WORD_EQ: "(All::(nat => bool) => bool) |
658 |
(%n::nat. |
|
| 17644 | 659 |
(RES_FORALL::('a::type word => bool)
|
660 |
=> ('a::type word => bool) => bool)
|
|
661 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
662 |
(%w1::'a::type word. |
|
663 |
(RES_FORALL::('a::type word => bool)
|
|
664 |
=> ('a::type word => bool) => bool)
|
|
665 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
666 |
(%w2::'a::type word. |
|
| 14847 | 667 |
(op -->::bool => bool => bool) |
668 |
((All::(nat => bool) => bool) |
|
669 |
(%k::nat. |
|
670 |
(op -->::bool => bool => bool) |
|
671 |
((op <::nat => nat => bool) k n) |
|
| 17644 | 672 |
((op =::'a::type => 'a::type => bool) |
673 |
((bit::nat => 'a::type word => 'a::type) k w1) |
|
674 |
((bit::nat => 'a::type word => 'a::type) k w2)))) |
|
675 |
((op =::'a::type word => 'a::type word => bool) w1 w2))))" |
|
| 14516 | 676 |
by (import word_base BIT_EQ_IMP_WORD_EQ) |
677 |
||
678 |
;end_setup |
|
679 |
||
680 |
;setup_theory word_num |
|
681 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
682 |
definition LVAL :: "('a => nat) => nat => 'a list => nat" where
|
| 17644 | 683 |
"LVAL == |
| 17652 | 684 |
%(f::'a::type => nat) b::nat. foldl (%(e::nat) x::'a::type. b * e + f x) 0" |
| 14516 | 685 |
|
| 17644 | 686 |
lemma LVAL_DEF: "ALL (f::'a::type => nat) (b::nat) l::'a::type list. |
| 17652 | 687 |
LVAL f b l = foldl (%(e::nat) x::'a::type. b * e + f x) 0 l" |
| 14516 | 688 |
by (import word_num LVAL_DEF) |
689 |
||
690 |
consts |
|
| 17652 | 691 |
NVAL :: "('a => nat) => nat => 'a word => nat"
|
| 14516 | 692 |
|
| 17644 | 693 |
specification (NVAL) NVAL_DEF: "ALL (f::'a::type => nat) (b::nat) l::'a::type list. |
694 |
NVAL f b (WORD l) = LVAL f b l" |
|
| 14516 | 695 |
by (import word_num NVAL_DEF) |
696 |
||
| 17652 | 697 |
lemma LVAL: "(ALL (x::'a::type => nat) xa::nat. LVAL x xa [] = 0) & |
| 17644 | 698 |
(ALL (x::'a::type list) (xa::'a::type => nat) (xb::nat) xc::'a::type. |
| 14516 | 699 |
LVAL xa xb (xc # x) = xa xc * xb ^ length x + LVAL xa xb x)" |
700 |
by (import word_num LVAL) |
|
701 |
||
| 17644 | 702 |
lemma LVAL_SNOC: "ALL (l::'a::type list) (h::'a::type) (f::'a::type => nat) b::nat. |
703 |
LVAL f b (SNOC h l) = LVAL f b l * b + f h" |
|
| 14516 | 704 |
by (import word_num LVAL_SNOC) |
705 |
||
| 17644 | 706 |
lemma LVAL_MAX: "ALL (l::'a::type list) (f::'a::type => nat) b::nat. |
707 |
(ALL x::'a::type. f x < b) --> LVAL f b l < b ^ length l" |
|
| 14516 | 708 |
by (import word_num LVAL_MAX) |
709 |
||
| 17644 | 710 |
lemma NVAL_MAX: "ALL (f::'a::type => nat) b::nat. |
711 |
(ALL x::'a::type. f x < b) --> |
|
712 |
(ALL n::nat. |
|
713 |
RES_FORALL (PWORDLEN n) (%w::'a::type word. NVAL f b w < b ^ n))" |
|
| 14516 | 714 |
by (import word_num NVAL_MAX) |
715 |
||
| 17652 | 716 |
lemma NVAL0: "ALL (x::'a::type => nat) xa::nat. NVAL x xa (WORD []) = 0" |
| 14516 | 717 |
by (import word_num NVAL0) |
718 |
||
| 17644 | 719 |
lemma NVAL1: "ALL (x::'a::type => nat) (xa::nat) xb::'a::type. |
720 |
NVAL x xa (WORD [xb]) = x xb" |
|
| 14516 | 721 |
by (import word_num NVAL1) |
722 |
||
| 17652 | 723 |
lemma NVAL_WORDLEN_0: "RES_FORALL (PWORDLEN 0) |
724 |
(%w::'a::type word. ALL (fv::'a::type => nat) r::nat. NVAL fv r w = 0)" |
|
| 14516 | 725 |
by (import word_num NVAL_WORDLEN_0) |
726 |
||
| 17644 | 727 |
lemma NVAL_WCAT1: "ALL (w::'a::type word) (f::'a::type => nat) (b::nat) x::'a::type. |
728 |
NVAL f b (WCAT (w, WORD [x])) = NVAL f b w * b + f x" |
|
| 14516 | 729 |
by (import word_num NVAL_WCAT1) |
730 |
||
| 17644 | 731 |
lemma NVAL_WCAT2: "ALL n::nat. |
| 14516 | 732 |
RES_FORALL (PWORDLEN n) |
| 17644 | 733 |
(%w::'a::type word. |
734 |
ALL (f::'a::type => nat) (b::nat) x::'a::type. |
|
735 |
NVAL f b (WCAT (WORD [x], w)) = f x * b ^ n + NVAL f b w)" |
|
| 14516 | 736 |
by (import word_num NVAL_WCAT2) |
737 |
||
| 17644 | 738 |
lemma NVAL_WCAT: "ALL (n::nat) m::nat. |
| 14516 | 739 |
RES_FORALL (PWORDLEN n) |
| 17644 | 740 |
(%w1::'a::type word. |
741 |
RES_FORALL (PWORDLEN m) |
|
742 |
(%w2::'a::type word. |
|
743 |
ALL (f::'a::type => nat) b::nat. |
|
744 |
NVAL f b (WCAT (w1, w2)) = |
|
745 |
NVAL f b w1 * b ^ m + NVAL f b w2))" |
|
| 14516 | 746 |
by (import word_num NVAL_WCAT) |
747 |
||
748 |
consts |
|
| 17652 | 749 |
NLIST :: "nat => (nat => 'a) => nat => nat => 'a list" |
| 14516 | 750 |
|
| 17652 | 751 |
specification (NLIST) NLIST_DEF: "(ALL (frep::nat => 'a::type) (b::nat) m::nat. NLIST 0 frep b m = []) & |
| 17644 | 752 |
(ALL (n::nat) (frep::nat => 'a::type) (b::nat) m::nat. |
| 14516 | 753 |
NLIST (Suc n) frep b m = |
754 |
SNOC (frep (m mod b)) (NLIST n frep b (m div b)))" |
|
755 |
by (import word_num NLIST_DEF) |
|
756 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
757 |
definition NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" where |
| 17644 | 758 |
"NWORD == |
759 |
%(n::nat) (frep::nat => 'a::type) (b::nat) m::nat. WORD (NLIST n frep b m)" |
|
| 14516 | 760 |
|
| 17644 | 761 |
lemma NWORD_DEF: "ALL (n::nat) (frep::nat => 'a::type) (b::nat) m::nat. |
762 |
NWORD n frep b m = WORD (NLIST n frep b m)" |
|
| 14516 | 763 |
by (import word_num NWORD_DEF) |
764 |
||
| 17644 | 765 |
lemma NWORD_LENGTH: "ALL (x::nat) (xa::nat => 'a::type) (xb::nat) xc::nat. |
766 |
WORDLEN (NWORD x xa xb xc) = x" |
|
| 14516 | 767 |
by (import word_num NWORD_LENGTH) |
768 |
||
| 17644 | 769 |
lemma NWORD_PWORDLEN: "ALL (x::nat) (xa::nat => 'a::type) (xb::nat) xc::nat. |
770 |
IN (NWORD x xa xb xc) (PWORDLEN x)" |
|
| 14516 | 771 |
by (import word_num NWORD_PWORDLEN) |
772 |
||
773 |
;end_setup |
|
774 |
||
775 |
;setup_theory word_bitop |
|
776 |
||
777 |
consts |
|
| 17652 | 778 |
PBITOP :: "('a word => 'b word) => bool"
|
| 14516 | 779 |
|
780 |
defs |
|
781 |
PBITOP_primdef: "PBITOP == |
|
782 |
GSPEC |
|
| 17644 | 783 |
(%oper::'a::type word => 'b::type word. |
| 14516 | 784 |
(oper, |
| 17644 | 785 |
ALL n::nat. |
| 14516 | 786 |
RES_FORALL (PWORDLEN n) |
| 17644 | 787 |
(%w::'a::type word. |
788 |
IN (oper w) (PWORDLEN n) & |
|
789 |
(ALL (m::nat) k::nat. |
|
790 |
m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))" |
|
| 14516 | 791 |
|
792 |
lemma PBITOP_def: "PBITOP = |
|
793 |
GSPEC |
|
| 17644 | 794 |
(%oper::'a::type word => 'b::type word. |
| 14516 | 795 |
(oper, |
| 17644 | 796 |
ALL n::nat. |
| 14516 | 797 |
RES_FORALL (PWORDLEN n) |
| 17644 | 798 |
(%w::'a::type word. |
799 |
IN (oper w) (PWORDLEN n) & |
|
800 |
(ALL (m::nat) k::nat. |
|
801 |
m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))" |
|
| 14516 | 802 |
by (import word_bitop PBITOP_def) |
803 |
||
| 17644 | 804 |
lemma IN_PBITOP: "ALL oper::'a::type word => 'b::type word. |
| 14516 | 805 |
IN oper PBITOP = |
| 17644 | 806 |
(ALL n::nat. |
| 14516 | 807 |
RES_FORALL (PWORDLEN n) |
| 17644 | 808 |
(%w::'a::type word. |
809 |
IN (oper w) (PWORDLEN n) & |
|
810 |
(ALL (m::nat) k::nat. |
|
811 |
m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w))))" |
|
| 14516 | 812 |
by (import word_bitop IN_PBITOP) |
813 |
||
814 |
lemma PBITOP_PWORDLEN: "RES_FORALL PBITOP |
|
| 17644 | 815 |
(%oper::'a::type word => 'b::type word. |
816 |
ALL n::nat. |
|
817 |
RES_FORALL (PWORDLEN n) |
|
818 |
(%w::'a::type word. IN (oper w) (PWORDLEN n)))" |
|
| 14516 | 819 |
by (import word_bitop PBITOP_PWORDLEN) |
820 |
||
821 |
lemma PBITOP_WSEG: "RES_FORALL PBITOP |
|
| 17644 | 822 |
(%oper::'a::type word => 'b::type word. |
823 |
ALL n::nat. |
|
| 14516 | 824 |
RES_FORALL (PWORDLEN n) |
| 17644 | 825 |
(%w::'a::type word. |
826 |
ALL (m::nat) k::nat. |
|
827 |
m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))" |
|
| 14516 | 828 |
by (import word_bitop PBITOP_WSEG) |
829 |
||
| 17644 | 830 |
lemma PBITOP_BIT: "(RES_FORALL::(('a::type word => 'b::type word) => bool)
|
831 |
=> (('a::type word => 'b::type word) => bool) => bool)
|
|
832 |
(PBITOP::('a::type word => 'b::type word) => bool)
|
|
833 |
(%oper::'a::type word => 'b::type word. |
|
| 14847 | 834 |
(All::(nat => bool) => bool) |
835 |
(%n::nat. |
|
| 17644 | 836 |
(RES_FORALL::('a::type word => bool)
|
837 |
=> ('a::type word => bool) => bool)
|
|
838 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
839 |
(%w::'a::type word. |
|
| 14847 | 840 |
(All::(nat => bool) => bool) |
841 |
(%k::nat. |
|
842 |
(op -->::bool => bool => bool) |
|
843 |
((op <::nat => nat => bool) k n) |
|
| 17644 | 844 |
((op =::'b::type word => 'b::type word => bool) |
| 14847 | 845 |
(oper |
| 17644 | 846 |
((WORD::'a::type list => 'a::type word) |
847 |
((op #::'a::type |
|
848 |
=> 'a::type list => 'a::type list) |
|
849 |
((bit::nat => 'a::type word => 'a::type) k w) |
|
850 |
([]::'a::type list)))) |
|
851 |
((WORD::'b::type list => 'b::type word) |
|
852 |
((op #::'b::type => 'b::type list => 'b::type list) |
|
853 |
((bit::nat => 'b::type word => 'b::type) k |
|
854 |
(oper w)) |
|
855 |
([]::'b::type list))))))))" |
|
| 14516 | 856 |
by (import word_bitop PBITOP_BIT) |
857 |
||
858 |
consts |
|
| 17652 | 859 |
PBITBOP :: "('a word => 'b word => 'c word) => bool"
|
| 14516 | 860 |
|
861 |
defs |
|
862 |
PBITBOP_primdef: "PBITBOP == |
|
863 |
GSPEC |
|
| 17644 | 864 |
(%oper::'a::type word => 'b::type word => 'c::type word. |
| 14516 | 865 |
(oper, |
| 17644 | 866 |
ALL n::nat. |
| 14516 | 867 |
RES_FORALL (PWORDLEN n) |
| 17644 | 868 |
(%w1::'a::type word. |
869 |
RES_FORALL (PWORDLEN n) |
|
870 |
(%w2::'b::type word. |
|
871 |
IN (oper w1 w2) (PWORDLEN n) & |
|
872 |
(ALL (m::nat) k::nat. |
|
873 |
m + k <= n --> |
|
874 |
oper (WSEG m k w1) (WSEG m k w2) = |
|
875 |
WSEG m k (oper w1 w2))))))" |
|
| 14516 | 876 |
|
877 |
lemma PBITBOP_def: "PBITBOP = |
|
878 |
GSPEC |
|
| 17644 | 879 |
(%oper::'a::type word => 'b::type word => 'c::type word. |
| 14516 | 880 |
(oper, |
| 17644 | 881 |
ALL n::nat. |
| 14516 | 882 |
RES_FORALL (PWORDLEN n) |
| 17644 | 883 |
(%w1::'a::type word. |
884 |
RES_FORALL (PWORDLEN n) |
|
885 |
(%w2::'b::type word. |
|
886 |
IN (oper w1 w2) (PWORDLEN n) & |
|
887 |
(ALL (m::nat) k::nat. |
|
888 |
m + k <= n --> |
|
889 |
oper (WSEG m k w1) (WSEG m k w2) = |
|
890 |
WSEG m k (oper w1 w2))))))" |
|
| 14516 | 891 |
by (import word_bitop PBITBOP_def) |
892 |
||
| 17644 | 893 |
lemma IN_PBITBOP: "ALL oper::'a::type word => 'b::type word => 'c::type word. |
| 14516 | 894 |
IN oper PBITBOP = |
| 17644 | 895 |
(ALL n::nat. |
| 14516 | 896 |
RES_FORALL (PWORDLEN n) |
| 17644 | 897 |
(%w1::'a::type word. |
898 |
RES_FORALL (PWORDLEN n) |
|
899 |
(%w2::'b::type word. |
|
900 |
IN (oper w1 w2) (PWORDLEN n) & |
|
901 |
(ALL (m::nat) k::nat. |
|
902 |
m + k <= n --> |
|
903 |
oper (WSEG m k w1) (WSEG m k w2) = |
|
904 |
WSEG m k (oper w1 w2)))))" |
|
| 14516 | 905 |
by (import word_bitop IN_PBITBOP) |
906 |
||
907 |
lemma PBITBOP_PWORDLEN: "RES_FORALL PBITBOP |
|
| 17644 | 908 |
(%oper::'a::type word => 'b::type word => 'c::type word. |
909 |
ALL n::nat. |
|
| 14516 | 910 |
RES_FORALL (PWORDLEN n) |
| 17644 | 911 |
(%w1::'a::type word. |
912 |
RES_FORALL (PWORDLEN n) |
|
913 |
(%w2::'b::type word. IN (oper w1 w2) (PWORDLEN n))))" |
|
| 14516 | 914 |
by (import word_bitop PBITBOP_PWORDLEN) |
915 |
||
916 |
lemma PBITBOP_WSEG: "RES_FORALL PBITBOP |
|
| 17644 | 917 |
(%oper::'a::type word => 'b::type word => 'c::type word. |
918 |
ALL n::nat. |
|
| 14516 | 919 |
RES_FORALL (PWORDLEN n) |
| 17644 | 920 |
(%w1::'a::type word. |
921 |
RES_FORALL (PWORDLEN n) |
|
922 |
(%w2::'b::type word. |
|
923 |
ALL (m::nat) k::nat. |
|
924 |
m + k <= n --> |
|
925 |
oper (WSEG m k w1) (WSEG m k w2) = |
|
926 |
WSEG m k (oper w1 w2))))" |
|
| 14516 | 927 |
by (import word_bitop PBITBOP_WSEG) |
928 |
||
| 17644 | 929 |
lemma PBITBOP_EXISTS: "ALL f::'a::type => 'b::type => 'c::type. |
930 |
EX x::'a::type word => 'b::type word => 'c::type word. |
|
931 |
ALL (l1::'a::type list) l2::'b::type list. |
|
932 |
x (WORD l1) (WORD l2) = WORD (map2 f l1 l2)" |
|
| 14516 | 933 |
by (import word_bitop PBITBOP_EXISTS) |
934 |
||
935 |
consts |
|
| 17652 | 936 |
WMAP :: "('a => 'b) => 'a word => 'b word"
|
| 14516 | 937 |
|
| 17644 | 938 |
specification (WMAP) WMAP_DEF: "ALL (f::'a::type => 'b::type) l::'a::type list. |
939 |
WMAP f (WORD l) = WORD (map f l)" |
|
| 14516 | 940 |
by (import word_bitop WMAP_DEF) |
941 |
||
| 17644 | 942 |
lemma WMAP_PWORDLEN: "RES_FORALL (PWORDLEN (n::nat)) |
943 |
(%w::'a::type word. |
|
944 |
ALL f::'a::type => 'b::type. IN (WMAP f w) (PWORDLEN n))" |
|
| 14516 | 945 |
by (import word_bitop WMAP_PWORDLEN) |
946 |
||
| 17644 | 947 |
lemma WMAP_0: "ALL x::'a::type => 'b::type. WMAP x (WORD []) = WORD []" |
| 14516 | 948 |
by (import word_bitop WMAP_0) |
949 |
||
| 14847 | 950 |
lemma WMAP_BIT: "(All::(nat => bool) => bool) |
951 |
(%n::nat. |
|
| 17644 | 952 |
(RES_FORALL::('a::type word => bool)
|
953 |
=> ('a::type word => bool) => bool)
|
|
954 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
955 |
(%w::'a::type word. |
|
| 14847 | 956 |
(All::(nat => bool) => bool) |
957 |
(%k::nat. |
|
958 |
(op -->::bool => bool => bool) |
|
959 |
((op <::nat => nat => bool) k n) |
|
| 17644 | 960 |
((All::(('a::type => 'b::type) => bool) => bool)
|
961 |
(%f::'a::type => 'b::type. |
|
962 |
(op =::'b::type => 'b::type => bool) |
|
963 |
((bit::nat => 'b::type word => 'b::type) k |
|
964 |
((WMAP::('a::type => 'b::type)
|
|
965 |
=> 'a::type word => 'b::type word) |
|
966 |
f w)) |
|
967 |
(f ((bit::nat => 'a::type word => 'a::type) k |
|
968 |
w)))))))" |
|
| 14516 | 969 |
by (import word_bitop WMAP_BIT) |
970 |
||
| 17644 | 971 |
lemma WMAP_WSEG: "ALL n::nat. |
| 14516 | 972 |
RES_FORALL (PWORDLEN n) |
| 17644 | 973 |
(%w::'a::type word. |
974 |
ALL (m::nat) k::nat. |
|
975 |
m + k <= n --> |
|
976 |
(ALL f::'a::type => 'b::type. |
|
977 |
WMAP f (WSEG m k w) = WSEG m k (WMAP f w)))" |
|
| 14516 | 978 |
by (import word_bitop WMAP_WSEG) |
979 |
||
| 17644 | 980 |
lemma WMAP_PBITOP: "ALL f::'a::type => 'b::type. IN (WMAP f) PBITOP" |
| 14516 | 981 |
by (import word_bitop WMAP_PBITOP) |
982 |
||
| 17644 | 983 |
lemma WMAP_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) f::'a::type => 'b::type. |
984 |
WMAP f (WCAT (w1, w2)) = WCAT (WMAP f w1, WMAP f w2)" |
|
| 14516 | 985 |
by (import word_bitop WMAP_WCAT) |
986 |
||
| 17644 | 987 |
lemma WMAP_o: "ALL (w::'a::type word) (f::'a::type => 'b::type) g::'b::type => 'c::type. |
988 |
WMAP g (WMAP f w) = WMAP (g o f) w" |
|
| 14516 | 989 |
by (import word_bitop WMAP_o) |
990 |
||
991 |
consts |
|
| 17652 | 992 |
FORALLBITS :: "('a => bool) => 'a word => bool"
|
| 14516 | 993 |
|
| 17644 | 994 |
specification (FORALLBITS) FORALLBITS_DEF: "ALL (P::'a::type => bool) l::'a::type list. |
995 |
FORALLBITS P (WORD l) = list_all P l" |
|
| 14516 | 996 |
by (import word_bitop FORALLBITS_DEF) |
997 |
||
| 14847 | 998 |
lemma FORALLBITS: "(All::(nat => bool) => bool) |
999 |
(%n::nat. |
|
| 17644 | 1000 |
(RES_FORALL::('a::type word => bool)
|
1001 |
=> ('a::type word => bool) => bool)
|
|
1002 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
1003 |
(%w::'a::type word. |
|
1004 |
(All::(('a::type => bool) => bool) => bool)
|
|
1005 |
(%P::'a::type => bool. |
|
| 14847 | 1006 |
(op =::bool => bool => bool) |
| 17644 | 1007 |
((FORALLBITS::('a::type => bool) => 'a::type word => bool) P
|
1008 |
w) |
|
| 14847 | 1009 |
((All::(nat => bool) => bool) |
1010 |
(%k::nat. |
|
1011 |
(op -->::bool => bool => bool) |
|
1012 |
((op <::nat => nat => bool) k n) |
|
| 17644 | 1013 |
(P ((bit::nat => 'a::type word => 'a::type) k |
1014 |
w)))))))" |
|
| 14516 | 1015 |
by (import word_bitop FORALLBITS) |
1016 |
||
| 17644 | 1017 |
lemma FORALLBITS_WSEG: "ALL n::nat. |
| 14516 | 1018 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1019 |
(%w::'a::type word. |
1020 |
ALL P::'a::type => bool. |
|
1021 |
FORALLBITS P w --> |
|
1022 |
(ALL (m::nat) k::nat. m + k <= n --> FORALLBITS P (WSEG m k w)))" |
|
| 14516 | 1023 |
by (import word_bitop FORALLBITS_WSEG) |
1024 |
||
| 17644 | 1025 |
lemma FORALLBITS_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) P::'a::type => bool. |
| 14516 | 1026 |
FORALLBITS P (WCAT (w1, w2)) = (FORALLBITS P w1 & FORALLBITS P w2)" |
1027 |
by (import word_bitop FORALLBITS_WCAT) |
|
1028 |
||
1029 |
consts |
|
| 17652 | 1030 |
EXISTSABIT :: "('a => bool) => 'a word => bool"
|
| 14516 | 1031 |
|
| 17644 | 1032 |
specification (EXISTSABIT) EXISTSABIT_DEF: "ALL (P::'a::type => bool) l::'a::type list. |
1033 |
EXISTSABIT P (WORD l) = list_exists P l" |
|
| 14516 | 1034 |
by (import word_bitop EXISTSABIT_DEF) |
1035 |
||
| 17644 | 1036 |
lemma NOT_EXISTSABIT: "ALL (P::'a::type => bool) w::'a::type word. |
1037 |
(~ EXISTSABIT P w) = FORALLBITS (Not o P) w" |
|
| 14516 | 1038 |
by (import word_bitop NOT_EXISTSABIT) |
1039 |
||
| 17644 | 1040 |
lemma NOT_FORALLBITS: "ALL (P::'a::type => bool) w::'a::type word. |
1041 |
(~ FORALLBITS P w) = EXISTSABIT (Not o P) w" |
|
| 14516 | 1042 |
by (import word_bitop NOT_FORALLBITS) |
1043 |
||
| 14847 | 1044 |
lemma EXISTSABIT: "(All::(nat => bool) => bool) |
1045 |
(%n::nat. |
|
| 17644 | 1046 |
(RES_FORALL::('a::type word => bool)
|
1047 |
=> ('a::type word => bool) => bool)
|
|
1048 |
((PWORDLEN::nat => 'a::type word => bool) n) |
|
1049 |
(%w::'a::type word. |
|
1050 |
(All::(('a::type => bool) => bool) => bool)
|
|
1051 |
(%P::'a::type => bool. |
|
| 14847 | 1052 |
(op =::bool => bool => bool) |
| 17644 | 1053 |
((EXISTSABIT::('a::type => bool) => 'a::type word => bool) P
|
1054 |
w) |
|
| 14847 | 1055 |
((Ex::(nat => bool) => bool) |
1056 |
(%k::nat. |
|
1057 |
(op &::bool => bool => bool) |
|
1058 |
((op <::nat => nat => bool) k n) |
|
| 17644 | 1059 |
(P ((bit::nat => 'a::type word => 'a::type) k |
1060 |
w)))))))" |
|
| 14516 | 1061 |
by (import word_bitop EXISTSABIT) |
1062 |
||
| 17644 | 1063 |
lemma EXISTSABIT_WSEG: "ALL n::nat. |
| 14516 | 1064 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1065 |
(%w::'a::type word. |
1066 |
ALL (m::nat) k::nat. |
|
1067 |
m + k <= n --> |
|
1068 |
(ALL P::'a::type => bool. |
|
1069 |
EXISTSABIT P (WSEG m k w) --> EXISTSABIT P w))" |
|
| 14516 | 1070 |
by (import word_bitop EXISTSABIT_WSEG) |
1071 |
||
| 17644 | 1072 |
lemma EXISTSABIT_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) P::'a::type => bool. |
| 14516 | 1073 |
EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)" |
1074 |
by (import word_bitop EXISTSABIT_WCAT) |
|
1075 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
1076 |
definition SHR :: "bool => 'a => 'a word => 'a word * 'a" where |
| 14516 | 1077 |
"SHR == |
| 17644 | 1078 |
%(f::bool) (b::'a::type) w::'a::type word. |
| 14516 | 1079 |
(WCAT |
| 17652 | 1080 |
(if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b], |
1081 |
WSEG (PRE (WORDLEN w)) 1 w), |
|
1082 |
bit 0 w)" |
|
| 14516 | 1083 |
|
| 17644 | 1084 |
lemma SHR_DEF: "ALL (f::bool) (b::'a::type) w::'a::type word. |
| 14516 | 1085 |
SHR f b w = |
1086 |
(WCAT |
|
| 17652 | 1087 |
(if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b], |
1088 |
WSEG (PRE (WORDLEN w)) 1 w), |
|
1089 |
bit 0 w)" |
|
| 14516 | 1090 |
by (import word_bitop SHR_DEF) |
1091 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
1092 |
definition SHL :: "bool => 'a word => 'a => 'a * 'a word" where |
| 14516 | 1093 |
"SHL == |
| 17644 | 1094 |
%(f::bool) (w::'a::type word) b::'a::type. |
| 14516 | 1095 |
(bit (PRE (WORDLEN w)) w, |
| 17652 | 1096 |
WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))" |
| 14516 | 1097 |
|
| 17644 | 1098 |
lemma SHL_DEF: "ALL (f::bool) (w::'a::type word) b::'a::type. |
| 14516 | 1099 |
SHL f w b = |
1100 |
(bit (PRE (WORDLEN w)) w, |
|
| 17652 | 1101 |
WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))" |
| 14516 | 1102 |
by (import word_bitop SHL_DEF) |
1103 |
||
| 17644 | 1104 |
lemma SHR_WSEG: "ALL n::nat. |
| 14516 | 1105 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1106 |
(%w::'a::type word. |
1107 |
ALL (m::nat) k::nat. |
|
1108 |
m + k <= n --> |
|
| 17652 | 1109 |
0 < m --> |
| 17644 | 1110 |
(ALL (f::bool) b::'a::type. |
1111 |
SHR f b (WSEG m k w) = |
|
1112 |
(if f |
|
| 17652 | 1113 |
then WCAT (WSEG 1 (k + (m - 1)) w, WSEG (m - 1) (k + 1) w) |
1114 |
else WCAT (WORD [b], WSEG (m - 1) (k + 1) w), |
|
| 17644 | 1115 |
bit k w)))" |
| 14516 | 1116 |
by (import word_bitop SHR_WSEG) |
1117 |
||
| 17644 | 1118 |
lemma SHR_WSEG_1F: "ALL n::nat. |
| 14516 | 1119 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1120 |
(%w::'a::type word. |
1121 |
ALL (b::'a::type) (m::nat) k::nat. |
|
1122 |
m + k <= n --> |
|
| 17652 | 1123 |
0 < m --> |
| 17644 | 1124 |
SHR False b (WSEG m k w) = |
| 17652 | 1125 |
(WCAT (WORD [b], WSEG (m - 1) (k + 1) w), bit k w))" |
| 14516 | 1126 |
by (import word_bitop SHR_WSEG_1F) |
1127 |
||
| 17644 | 1128 |
lemma SHR_WSEG_NF: "ALL n::nat. |
| 14516 | 1129 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1130 |
(%w::'a::type word. |
1131 |
ALL (m::nat) k::nat. |
|
1132 |
m + k < n --> |
|
| 17652 | 1133 |
0 < m --> |
| 17644 | 1134 |
SHR False (bit (m + k) w) (WSEG m k w) = |
| 17652 | 1135 |
(WSEG m (k + 1) w, bit k w))" |
| 14516 | 1136 |
by (import word_bitop SHR_WSEG_NF) |
1137 |
||
| 17644 | 1138 |
lemma SHL_WSEG: "ALL n::nat. |
| 14516 | 1139 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1140 |
(%w::'a::type word. |
1141 |
ALL (m::nat) k::nat. |
|
1142 |
m + k <= n --> |
|
| 17652 | 1143 |
0 < m --> |
| 17644 | 1144 |
(ALL (f::bool) b::'a::type. |
1145 |
SHL f (WSEG m k w) b = |
|
| 17652 | 1146 |
(bit (k + (m - 1)) w, |
1147 |
if f then WCAT (WSEG (m - 1) k w, WSEG 1 k w) |
|
1148 |
else WCAT (WSEG (m - 1) k w, WORD [b]))))" |
|
| 14516 | 1149 |
by (import word_bitop SHL_WSEG) |
1150 |
||
| 17644 | 1151 |
lemma SHL_WSEG_1F: "ALL n::nat. |
| 14516 | 1152 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1153 |
(%w::'a::type word. |
1154 |
ALL (b::'a::type) (m::nat) k::nat. |
|
1155 |
m + k <= n --> |
|
| 17652 | 1156 |
0 < m --> |
| 17644 | 1157 |
SHL False (WSEG m k w) b = |
| 17652 | 1158 |
(bit (k + (m - 1)) w, WCAT (WSEG (m - 1) k w, WORD [b])))" |
| 14516 | 1159 |
by (import word_bitop SHL_WSEG_1F) |
1160 |
||
| 17644 | 1161 |
lemma SHL_WSEG_NF: "ALL n::nat. |
| 14516 | 1162 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1163 |
(%w::'a::type word. |
1164 |
ALL (m::nat) k::nat. |
|
1165 |
m + k <= n --> |
|
| 17652 | 1166 |
0 < m --> |
1167 |
0 < k --> |
|
1168 |
SHL False (WSEG m k w) (bit (k - 1) w) = |
|
1169 |
(bit (k + (m - 1)) w, WSEG m (k - 1) w))" |
|
| 14516 | 1170 |
by (import word_bitop SHL_WSEG_NF) |
1171 |
||
| 17644 | 1172 |
lemma WSEG_SHL: "ALL n::nat. |
| 14516 | 1173 |
RES_FORALL (PWORDLEN (Suc n)) |
| 17644 | 1174 |
(%w::'a::type word. |
1175 |
ALL (m::nat) k::nat. |
|
| 17652 | 1176 |
0 < k & m + k <= Suc n --> |
| 17644 | 1177 |
(ALL b::'a::type. |
| 17652 | 1178 |
WSEG m k (snd (SHL (f::bool) w b)) = WSEG m (k - 1) w))" |
| 14516 | 1179 |
by (import word_bitop WSEG_SHL) |
1180 |
||
| 17644 | 1181 |
lemma WSEG_SHL_0: "ALL n::nat. |
| 14516 | 1182 |
RES_FORALL (PWORDLEN (Suc n)) |
| 17644 | 1183 |
(%w::'a::type word. |
1184 |
ALL (m::nat) b::'a::type. |
|
| 17652 | 1185 |
0 < m & m <= Suc n --> |
1186 |
WSEG m 0 (snd (SHL (f::bool) w b)) = |
|
1187 |
WCAT (WSEG (m - 1) 0 w, if f then WSEG 1 0 w else WORD [b]))" |
|
| 14516 | 1188 |
by (import word_bitop WSEG_SHL_0) |
1189 |
||
1190 |
;end_setup |
|
1191 |
||
1192 |
;setup_theory bword_num |
|
1193 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
1194 |
definition BV :: "bool => nat" where |
| 17652 | 1195 |
"BV == %b::bool. if b then Suc 0 else 0" |
| 14516 | 1196 |
|
| 17652 | 1197 |
lemma BV_DEF: "ALL b::bool. BV b = (if b then Suc 0 else 0)" |
| 14516 | 1198 |
by (import bword_num BV_DEF) |
1199 |
||
1200 |
consts |
|
1201 |
BNVAL :: "bool word => nat" |
|
1202 |
||
| 17652 | 1203 |
specification (BNVAL) BNVAL_DEF: "ALL l::bool list. BNVAL (WORD l) = LVAL BV 2 l" |
| 14516 | 1204 |
by (import bword_num BNVAL_DEF) |
1205 |
||
| 17652 | 1206 |
lemma BV_LESS_2: "ALL x::bool. BV x < 2" |
| 14516 | 1207 |
by (import bword_num BV_LESS_2) |
1208 |
||
| 17652 | 1209 |
lemma BNVAL_NVAL: "ALL w::bool word. BNVAL w = NVAL BV 2 w" |
| 14516 | 1210 |
by (import bword_num BNVAL_NVAL) |
1211 |
||
| 17652 | 1212 |
lemma BNVAL0: "BNVAL (WORD []) = 0" |
| 14516 | 1213 |
by (import bword_num BNVAL0) |
1214 |
||
| 17644 | 1215 |
lemma BNVAL_11: "ALL (w1::bool word) w2::bool word. |
1216 |
WORDLEN w1 = WORDLEN w2 --> BNVAL w1 = BNVAL w2 --> w1 = w2" |
|
| 14516 | 1217 |
by (import bword_num BNVAL_11) |
1218 |
||
| 17644 | 1219 |
lemma BNVAL_ONTO: "ALL w::bool word. Ex (op = (BNVAL w))" |
| 14516 | 1220 |
by (import bword_num BNVAL_ONTO) |
1221 |
||
| 17652 | 1222 |
lemma BNVAL_MAX: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. BNVAL w < 2 ^ n)" |
| 14516 | 1223 |
by (import bword_num BNVAL_MAX) |
1224 |
||
| 17644 | 1225 |
lemma BNVAL_WCAT1: "ALL n::nat. |
| 14516 | 1226 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1227 |
(%w::bool word. |
| 17652 | 1228 |
ALL x::bool. BNVAL (WCAT (w, WORD [x])) = BNVAL w * 2 + BV x)" |
| 14516 | 1229 |
by (import bword_num BNVAL_WCAT1) |
1230 |
||
| 17644 | 1231 |
lemma BNVAL_WCAT2: "ALL n::nat. |
| 14516 | 1232 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1233 |
(%w::bool word. |
| 17652 | 1234 |
ALL x::bool. BNVAL (WCAT (WORD [x], w)) = BV x * 2 ^ n + BNVAL w)" |
| 14516 | 1235 |
by (import bword_num BNVAL_WCAT2) |
1236 |
||
| 17644 | 1237 |
lemma BNVAL_WCAT: "ALL (n::nat) m::nat. |
| 14516 | 1238 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1239 |
(%w1::bool word. |
1240 |
RES_FORALL (PWORDLEN m) |
|
1241 |
(%w2::bool word. |
|
| 17652 | 1242 |
BNVAL (WCAT (w1, w2)) = BNVAL w1 * 2 ^ m + BNVAL w2))" |
| 14516 | 1243 |
by (import bword_num BNVAL_WCAT) |
1244 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
1245 |
definition VB :: "nat => bool" where |
| 17652 | 1246 |
"VB == %n::nat. n mod 2 ~= 0" |
| 14516 | 1247 |
|
| 17652 | 1248 |
lemma VB_DEF: "ALL n::nat. VB n = (n mod 2 ~= 0)" |
| 14516 | 1249 |
by (import bword_num VB_DEF) |
1250 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
26086
diff
changeset
|
1251 |
definition NBWORD :: "nat => nat => bool word" where |
| 17652 | 1252 |
"NBWORD == %(n::nat) m::nat. WORD (NLIST n VB 2 m)" |
| 14516 | 1253 |
|
| 17652 | 1254 |
lemma NBWORD_DEF: "ALL (n::nat) m::nat. NBWORD n m = WORD (NLIST n VB 2 m)" |
| 14516 | 1255 |
by (import bword_num NBWORD_DEF) |
1256 |
||
| 17652 | 1257 |
lemma NBWORD0: "ALL x::nat. NBWORD 0 x = WORD []" |
| 14516 | 1258 |
by (import bword_num NBWORD0) |
1259 |
||
| 17644 | 1260 |
lemma WORDLEN_NBWORD: "ALL (x::nat) xa::nat. WORDLEN (NBWORD x xa) = x" |
| 14516 | 1261 |
by (import bword_num WORDLEN_NBWORD) |
1262 |
||
| 17644 | 1263 |
lemma PWORDLEN_NBWORD: "ALL (x::nat) xa::nat. IN (NBWORD x xa) (PWORDLEN x)" |
| 14516 | 1264 |
by (import bword_num PWORDLEN_NBWORD) |
1265 |
||
| 17644 | 1266 |
lemma NBWORD_SUC: "ALL (n::nat) m::nat. |
| 17652 | 1267 |
NBWORD (Suc n) m = WCAT (NBWORD n (m div 2), WORD [VB (m mod 2)])" |
| 14516 | 1268 |
by (import bword_num NBWORD_SUC) |
1269 |
||
| 17644 | 1270 |
lemma VB_BV: "ALL x::bool. VB (BV x) = x" |
| 14516 | 1271 |
by (import bword_num VB_BV) |
1272 |
||
| 14847 | 1273 |
lemma BV_VB: "(All::(nat => bool) => bool) |
1274 |
(%x::nat. |
|
1275 |
(op -->::bool => bool => bool) |
|
1276 |
((op <::nat => nat => bool) x |
|
| 20485 | 1277 |
((number_of \<Colon> int => nat) |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
1278 |
((Int.Bit0 \<Colon> int => int) |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
1279 |
((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))) |
| 14847 | 1280 |
((op =::nat => nat => bool) ((BV::bool => nat) ((VB::nat => bool) x)) |
1281 |
x))" |
|
| 14516 | 1282 |
by (import bword_num BV_VB) |
1283 |
||
| 17644 | 1284 |
lemma NBWORD_BNVAL: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. NBWORD n (BNVAL w) = w)" |
| 14516 | 1285 |
by (import bword_num NBWORD_BNVAL) |
1286 |
||
| 17652 | 1287 |
lemma BNVAL_NBWORD: "ALL (n::nat) m::nat. m < 2 ^ n --> BNVAL (NBWORD n m) = m" |
| 14516 | 1288 |
by (import bword_num BNVAL_NBWORD) |
1289 |
||
| 17644 | 1290 |
lemma ZERO_WORD_VAL: "RES_FORALL (PWORDLEN (n::nat)) |
| 17652 | 1291 |
(%w::bool word. (w = NBWORD n 0) = (BNVAL w = 0))" |
| 14516 | 1292 |
by (import bword_num ZERO_WORD_VAL) |
1293 |
||
| 17652 | 1294 |
lemma WCAT_NBWORD_0: "ALL (n1::nat) n2::nat. WCAT (NBWORD n1 0, NBWORD n2 0) = NBWORD (n1 + n2) 0" |
| 14516 | 1295 |
by (import bword_num WCAT_NBWORD_0) |
1296 |
||
| 17644 | 1297 |
lemma WSPLIT_NBWORD_0: "ALL (n::nat) m::nat. |
| 17652 | 1298 |
m <= n --> WSPLIT m (NBWORD n 0) = (NBWORD (n - m) 0, NBWORD m 0)" |
| 14516 | 1299 |
by (import bword_num WSPLIT_NBWORD_0) |
1300 |
||
| 14847 | 1301 |
lemma EQ_NBWORD0_SPLIT: "(All::(nat => bool) => bool) |
1302 |
(%n::nat. |
|
1303 |
(RES_FORALL::(bool word => bool) => (bool word => bool) => bool) |
|
1304 |
((PWORDLEN::nat => bool word => bool) n) |
|
1305 |
(%w::bool word. |
|
1306 |
(All::(nat => bool) => bool) |
|
1307 |
(%m::nat. |
|
1308 |
(op -->::bool => bool => bool) |
|
1309 |
((op <=::nat => nat => bool) m n) |
|
1310 |
((op =::bool => bool => bool) |
|
1311 |
((op =::bool word => bool word => bool) w |
|
1312 |
((NBWORD::nat => nat => bool word) n (0::nat))) |
|
1313 |
((op &::bool => bool => bool) |
|
1314 |
((op =::bool word => bool word => bool) |
|
1315 |
((WSEG::nat => nat => bool word => bool word) |
|
1316 |
((op -::nat => nat => nat) n m) m w) |
|
1317 |
((NBWORD::nat => nat => bool word) |
|
1318 |
((op -::nat => nat => nat) n m) (0::nat))) |
|
1319 |
((op =::bool word => bool word => bool) |
|
1320 |
((WSEG::nat => nat => bool word => bool word) m |
|
1321 |
(0::nat) w) |
|
1322 |
((NBWORD::nat => nat => bool word) m (0::nat))))))))" |
|
| 14516 | 1323 |
by (import bword_num EQ_NBWORD0_SPLIT) |
1324 |
||
| 17652 | 1325 |
lemma NBWORD_MOD: "ALL (n::nat) m::nat. NBWORD n (m mod 2 ^ n) = NBWORD n m" |
| 14516 | 1326 |
by (import bword_num NBWORD_MOD) |
1327 |
||
| 17652 | 1328 |
lemma WSEG_NBWORD_SUC: "ALL (n::nat) m::nat. WSEG n 0 (NBWORD (Suc n) m) = NBWORD n m" |
| 14516 | 1329 |
by (import bword_num WSEG_NBWORD_SUC) |
1330 |
||
| 17644 | 1331 |
lemma NBWORD_SUC_WSEG: "ALL n::nat. |
1332 |
RES_FORALL (PWORDLEN (Suc n)) |
|
| 17652 | 1333 |
(%w::bool word. NBWORD n (BNVAL w) = WSEG n 0 w)" |
| 14516 | 1334 |
by (import bword_num NBWORD_SUC_WSEG) |
1335 |
||
| 17652 | 1336 |
lemma DOUBL_EQ_SHL: "ALL x>0. |
| 14516 | 1337 |
RES_FORALL (PWORDLEN x) |
| 17644 | 1338 |
(%xa::bool word. |
1339 |
ALL xb::bool. |
|
1340 |
NBWORD x (BNVAL xa + BNVAL xa + BV xb) = snd (SHL False xa xb))" |
|
| 14516 | 1341 |
by (import bword_num DOUBL_EQ_SHL) |
1342 |
||
| 17652 | 1343 |
lemma MSB_NBWORD: "ALL (n::nat) m::nat. bit n (NBWORD (Suc n) m) = VB (m div 2 ^ n mod 2)" |
| 14516 | 1344 |
by (import bword_num MSB_NBWORD) |
1345 |
||
| 17644 | 1346 |
lemma NBWORD_SPLIT: "ALL (n1::nat) (n2::nat) m::nat. |
| 17652 | 1347 |
NBWORD (n1 + n2) m = WCAT (NBWORD n1 (m div 2 ^ n2), NBWORD n2 m)" |
| 14516 | 1348 |
by (import bword_num NBWORD_SPLIT) |
1349 |
||
| 17644 | 1350 |
lemma WSEG_NBWORD: "ALL (m::nat) (k::nat) n::nat. |
1351 |
m + k <= n --> |
|
| 17652 | 1352 |
(ALL l::nat. WSEG m k (NBWORD n l) = NBWORD m (l div 2 ^ k))" |
| 14516 | 1353 |
by (import bword_num WSEG_NBWORD) |
1354 |
||
| 17644 | 1355 |
lemma NBWORD_SUC_FST: "ALL (n::nat) x::nat. |
| 17652 | 1356 |
NBWORD (Suc n) x = WCAT (WORD [VB (x div 2 ^ n mod 2)], NBWORD n x)" |
| 14516 | 1357 |
by (import bword_num NBWORD_SUC_FST) |
1358 |
||
| 17652 | 1359 |
lemma BIT_NBWORD0: "ALL (k::nat) n::nat. k < n --> bit k (NBWORD n 0) = False" |
| 14516 | 1360 |
by (import bword_num BIT_NBWORD0) |
1361 |
||
| 17644 | 1362 |
lemma ADD_BNVAL_LEFT: "ALL n::nat. |
| 14516 | 1363 |
RES_FORALL (PWORDLEN (Suc n)) |
| 17644 | 1364 |
(%w1::bool word. |
1365 |
RES_FORALL (PWORDLEN (Suc n)) |
|
1366 |
(%w2::bool word. |
|
1367 |
BNVAL w1 + BNVAL w2 = |
|
| 17652 | 1368 |
(BV (bit n w1) + BV (bit n w2)) * 2 ^ n + |
1369 |
(BNVAL (WSEG n 0 w1) + BNVAL (WSEG n 0 w2))))" |
|
| 14516 | 1370 |
by (import bword_num ADD_BNVAL_LEFT) |
1371 |
||
| 17644 | 1372 |
lemma ADD_BNVAL_RIGHT: "ALL n::nat. |
| 14516 | 1373 |
RES_FORALL (PWORDLEN (Suc n)) |
| 17644 | 1374 |
(%w1::bool word. |
1375 |
RES_FORALL (PWORDLEN (Suc n)) |
|
1376 |
(%w2::bool word. |
|
1377 |
BNVAL w1 + BNVAL w2 = |
|
| 17652 | 1378 |
(BNVAL (WSEG n 1 w1) + BNVAL (WSEG n 1 w2)) * 2 + |
1379 |
(BV (bit 0 w1) + BV (bit 0 w2))))" |
|
| 14516 | 1380 |
by (import bword_num ADD_BNVAL_RIGHT) |
1381 |
||
| 17644 | 1382 |
lemma ADD_BNVAL_SPLIT: "ALL (n1::nat) n2::nat. |
| 14516 | 1383 |
RES_FORALL (PWORDLEN (n1 + n2)) |
| 17644 | 1384 |
(%w1::bool word. |
1385 |
RES_FORALL (PWORDLEN (n1 + n2)) |
|
1386 |
(%w2::bool word. |
|
1387 |
BNVAL w1 + BNVAL w2 = |
|
| 17652 | 1388 |
(BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2)) * 2 ^ n2 + |
1389 |
(BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2))))" |
|
| 14516 | 1390 |
by (import bword_num ADD_BNVAL_SPLIT) |
1391 |
||
1392 |
;end_setup |
|
1393 |
||
1394 |
;setup_theory bword_arith |
|
1395 |
||
1396 |
consts |
|
1397 |
ACARRY :: "nat => bool word => bool word => bool => bool" |
|
1398 |
||
| 17652 | 1399 |
specification (ACARRY) ACARRY_DEF: "(ALL (w1::bool word) (w2::bool word) cin::bool. ACARRY 0 w1 w2 cin = cin) & |
| 17644 | 1400 |
(ALL (n::nat) (w1::bool word) (w2::bool word) cin::bool. |
| 14516 | 1401 |
ACARRY (Suc n) w1 w2 cin = |
| 17652 | 1402 |
VB ((BV (bit n w1) + BV (bit n w2) + BV (ACARRY n w1 w2 cin)) div 2))" |
| 14516 | 1403 |
by (import bword_arith ACARRY_DEF) |
1404 |
||
1405 |
consts |
|
1406 |
ICARRY :: "nat => bool word => bool word => bool => bool" |
|
1407 |
||
| 17652 | 1408 |
specification (ICARRY) ICARRY_DEF: "(ALL (w1::bool word) (w2::bool word) cin::bool. ICARRY 0 w1 w2 cin = cin) & |
| 17644 | 1409 |
(ALL (n::nat) (w1::bool word) (w2::bool word) cin::bool. |
| 14516 | 1410 |
ICARRY (Suc n) w1 w2 cin = |
1411 |
(bit n w1 & bit n w2 | (bit n w1 | bit n w2) & ICARRY n w1 w2 cin))" |
|
1412 |
by (import bword_arith ICARRY_DEF) |
|
1413 |
||
| 17644 | 1414 |
lemma ACARRY_EQ_ICARRY: "ALL n::nat. |
| 14516 | 1415 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1416 |
(%w1::bool word. |
1417 |
RES_FORALL (PWORDLEN n) |
|
1418 |
(%w2::bool word. |
|
1419 |
ALL (cin::bool) k::nat. |
|
1420 |
k <= n --> ACARRY k w1 w2 cin = ICARRY k w1 w2 cin))" |
|
| 14516 | 1421 |
by (import bword_arith ACARRY_EQ_ICARRY) |
1422 |
||
| 17652 | 1423 |
lemma BNVAL_LESS_EQ: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. BNVAL w <= 2 ^ n - 1)" |
| 14516 | 1424 |
by (import bword_arith BNVAL_LESS_EQ) |
1425 |
||
| 17644 | 1426 |
lemma ADD_BNVAL_LESS_EQ1: "ALL (n::nat) cin::bool. |
| 14516 | 1427 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1428 |
(%w1::bool word. |
1429 |
RES_FORALL (PWORDLEN n) |
|
1430 |
(%w2::bool word. |
|
| 17652 | 1431 |
(BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n <= Suc 0))" |
| 14516 | 1432 |
by (import bword_arith ADD_BNVAL_LESS_EQ1) |
1433 |
||
| 17644 | 1434 |
lemma ADD_BV_BNVAL_DIV_LESS_EQ1: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool. |
| 14516 | 1435 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1436 |
(%w1::bool word. |
1437 |
RES_FORALL (PWORDLEN n) |
|
1438 |
(%w2::bool word. |
|
1439 |
(BV x1 + BV x2 + |
|
| 17652 | 1440 |
(BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n) div |
1441 |
2 |
|
1442 |
<= 1))" |
|
| 14516 | 1443 |
by (import bword_arith ADD_BV_BNVAL_DIV_LESS_EQ1) |
1444 |
||
| 17644 | 1445 |
lemma ADD_BV_BNVAL_LESS_EQ: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool. |
| 14516 | 1446 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1447 |
(%w1::bool word. |
1448 |
RES_FORALL (PWORDLEN n) |
|
1449 |
(%w2::bool word. |
|
1450 |
BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin)) |
|
| 17652 | 1451 |
<= Suc (2 ^ Suc n)))" |
| 14516 | 1452 |
by (import bword_arith ADD_BV_BNVAL_LESS_EQ) |
1453 |
||
| 17644 | 1454 |
lemma ADD_BV_BNVAL_LESS_EQ1: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool. |
| 14516 | 1455 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1456 |
(%w1::bool word. |
1457 |
RES_FORALL (PWORDLEN n) |
|
1458 |
(%w2::bool word. |
|
1459 |
(BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))) div |
|
| 17652 | 1460 |
2 ^ Suc n |
1461 |
<= 1))" |
|
| 14516 | 1462 |
by (import bword_arith ADD_BV_BNVAL_LESS_EQ1) |
1463 |
||
| 14847 | 1464 |
lemma ACARRY_EQ_ADD_DIV: "(All::(nat => bool) => bool) |
1465 |
(%n::nat. |
|
1466 |
(RES_FORALL::(bool word => bool) => (bool word => bool) => bool) |
|
1467 |
((PWORDLEN::nat => bool word => bool) n) |
|
1468 |
(%w1::bool word. |
|
1469 |
(RES_FORALL::(bool word => bool) => (bool word => bool) => bool) |
|
1470 |
((PWORDLEN::nat => bool word => bool) n) |
|
1471 |
(%w2::bool word. |
|
1472 |
(All::(nat => bool) => bool) |
|
1473 |
(%k::nat. |
|
1474 |
(op -->::bool => bool => bool) |
|
1475 |
((op <::nat => nat => bool) k n) |
|
1476 |
((op =::nat => nat => bool) |
|
1477 |
((BV::bool => nat) |
|
1478 |
((ACARRY::nat |
|
1479 |
=> bool word |
|
1480 |
=> bool word => bool => bool) |
|
1481 |
k w1 w2 (cin::bool))) |
|
1482 |
((op div::nat => nat => nat) |
|
1483 |
((op +::nat => nat => nat) |
|
1484 |
((op +::nat => nat => nat) |
|
1485 |
((BNVAL::bool word => nat) |
|
1486 |
((WSEG::nat => nat => bool word => bool word) |
|
1487 |
k (0::nat) w1)) |
|
1488 |
((BNVAL::bool word => nat) |
|
1489 |
((WSEG::nat => nat => bool word => bool word) |
|
1490 |
k (0::nat) w2))) |
|
1491 |
((BV::bool => nat) cin)) |
|
1492 |
((op ^::nat => nat => nat) |
|
| 20485 | 1493 |
((number_of \<Colon> int => nat) |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
1494 |
((Int.Bit0 \<Colon> int => int) |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
1495 |
((Int.Bit1 \<Colon> int => int) |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25919
diff
changeset
|
1496 |
(Int.Pls \<Colon> int)))) |
| 14847 | 1497 |
k)))))))" |
| 14516 | 1498 |
by (import bword_arith ACARRY_EQ_ADD_DIV) |
1499 |
||
| 17644 | 1500 |
lemma ADD_WORD_SPLIT: "ALL (n1::nat) n2::nat. |
| 14516 | 1501 |
RES_FORALL (PWORDLEN (n1 + n2)) |
| 17644 | 1502 |
(%w1::bool word. |
1503 |
RES_FORALL (PWORDLEN (n1 + n2)) |
|
1504 |
(%w2::bool word. |
|
1505 |
ALL cin::bool. |
|
1506 |
NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) = |
|
1507 |
WCAT |
|
1508 |
(NBWORD n1 |
|
1509 |
(BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) + |
|
1510 |
BV (ACARRY n2 w1 w2 cin)), |
|
1511 |
NBWORD n2 |
|
| 17652 | 1512 |
(BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + BV cin))))" |
| 14516 | 1513 |
by (import bword_arith ADD_WORD_SPLIT) |
1514 |
||
| 17644 | 1515 |
lemma WSEG_NBWORD_ADD: "ALL n::nat. |
| 14516 | 1516 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1517 |
(%w1::bool word. |
1518 |
RES_FORALL (PWORDLEN n) |
|
1519 |
(%w2::bool word. |
|
1520 |
ALL (m::nat) (k::nat) cin::bool. |
|
1521 |
m + k <= n --> |
|
1522 |
WSEG m k (NBWORD n (BNVAL w1 + BNVAL w2 + BV cin)) = |
|
1523 |
NBWORD m |
|
1524 |
(BNVAL (WSEG m k w1) + BNVAL (WSEG m k w2) + |
|
1525 |
BV (ACARRY k w1 w2 cin))))" |
|
| 14516 | 1526 |
by (import bword_arith WSEG_NBWORD_ADD) |
1527 |
||
| 17644 | 1528 |
lemma ADD_NBWORD_EQ0_SPLIT: "ALL (n1::nat) n2::nat. |
| 14516 | 1529 |
RES_FORALL (PWORDLEN (n1 + n2)) |
| 17644 | 1530 |
(%w1::bool word. |
1531 |
RES_FORALL (PWORDLEN (n1 + n2)) |
|
1532 |
(%w2::bool word. |
|
1533 |
ALL cin::bool. |
|
1534 |
(NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) = |
|
| 17652 | 1535 |
NBWORD (n1 + n2) 0) = |
| 17644 | 1536 |
(NBWORD n1 |
1537 |
(BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) + |
|
1538 |
BV (ACARRY n2 w1 w2 cin)) = |
|
| 17652 | 1539 |
NBWORD n1 0 & |
| 17644 | 1540 |
NBWORD n2 |
| 17652 | 1541 |
(BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + BV cin) = |
1542 |
NBWORD n2 0)))" |
|
| 14516 | 1543 |
by (import bword_arith ADD_NBWORD_EQ0_SPLIT) |
1544 |
||
| 17644 | 1545 |
lemma ACARRY_MSB: "ALL n::nat. |
| 14516 | 1546 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1547 |
(%w1::bool word. |
1548 |
RES_FORALL (PWORDLEN n) |
|
1549 |
(%w2::bool word. |
|
1550 |
ALL cin::bool. |
|
1551 |
ACARRY n w1 w2 cin = |
|
1552 |
bit n (NBWORD (Suc n) (BNVAL w1 + BNVAL w2 + BV cin))))" |
|
| 14516 | 1553 |
by (import bword_arith ACARRY_MSB) |
1554 |
||
| 17644 | 1555 |
lemma ACARRY_WSEG: "ALL n::nat. |
| 14516 | 1556 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1557 |
(%w1::bool word. |
1558 |
RES_FORALL (PWORDLEN n) |
|
1559 |
(%w2::bool word. |
|
1560 |
ALL (cin::bool) (k::nat) m::nat. |
|
1561 |
k < m & m <= n --> |
|
| 17652 | 1562 |
ACARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin = |
| 17644 | 1563 |
ACARRY k w1 w2 cin))" |
| 14516 | 1564 |
by (import bword_arith ACARRY_WSEG) |
1565 |
||
| 17644 | 1566 |
lemma ICARRY_WSEG: "ALL n::nat. |
| 14516 | 1567 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1568 |
(%w1::bool word. |
1569 |
RES_FORALL (PWORDLEN n) |
|
1570 |
(%w2::bool word. |
|
1571 |
ALL (cin::bool) (k::nat) m::nat. |
|
1572 |
k < m & m <= n --> |
|
| 17652 | 1573 |
ICARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin = |
| 17644 | 1574 |
ICARRY k w1 w2 cin))" |
| 14516 | 1575 |
by (import bword_arith ICARRY_WSEG) |
1576 |
||
| 17644 | 1577 |
lemma ACARRY_ACARRY_WSEG: "ALL n::nat. |
| 14516 | 1578 |
RES_FORALL (PWORDLEN n) |
| 17644 | 1579 |
(%w1::bool word. |
1580 |
RES_FORALL (PWORDLEN n) |
|
1581 |
(%w2::bool word. |
|
1582 |
ALL (cin::bool) (m::nat) (k1::nat) k2::nat. |
|
1583 |
k1 < m & k2 < n & m + k2 <= n --> |
|
1584 |
ACARRY k1 (WSEG m k2 w1) (WSEG m k2 w2) |
|
1585 |
(ACARRY k2 w1 w2 cin) = |
|
1586 |
ACARRY (k1 + k2) w1 w2 cin))" |
|
| 14516 | 1587 |
by (import bword_arith ACARRY_ACARRY_WSEG) |
1588 |
||
1589 |
;end_setup |
|
1590 |
||
1591 |
;setup_theory bword_bitop |
|
1592 |
||
1593 |
consts |
|
1594 |
WNOT :: "bool word => bool word" |
|
1595 |
||
| 17644 | 1596 |
specification (WNOT) WNOT_DEF: "ALL l::bool list. WNOT (WORD l) = WORD (map Not l)" |
| 14516 | 1597 |
by (import bword_bitop WNOT_DEF) |
1598 |
||
1599 |
lemma PBITOP_WNOT: "IN WNOT PBITOP" |
|
1600 |
by (import bword_bitop PBITOP_WNOT) |
|
1601 |
||
| 17644 | 1602 |
lemma WNOT_WNOT: "ALL w::bool word. WNOT (WNOT w) = w" |
| 14516 | 1603 |
by (import bword_bitop WNOT_WNOT) |
1604 |
||
| 17644 | 1605 |
lemma WCAT_WNOT: "ALL (n1::nat) n2::nat. |
| 14516 | 1606 |
RES_FORALL (PWORDLEN n1) |
| 17644 | 1607 |
(%w1::bool word. |
1608 |
RES_FORALL (PWORDLEN n2) |
|
1609 |
(%w2::bool word. WCAT (WNOT w1, WNOT w2) = WNOT (WCAT (w1, w2))))" |
|
| 14516 | 1610 |
by (import bword_bitop WCAT_WNOT) |
1611 |
||
1612 |
consts |
|
1613 |
WAND :: "bool word => bool word => bool word" |
|
1614 |
||
| 17644 | 1615 |
specification (WAND) WAND_DEF: "ALL (l1::bool list) l2::bool list. |
1616 |
WAND (WORD l1) (WORD l2) = WORD (map2 op & l1 l2)" |
|
| 14516 | 1617 |
by (import bword_bitop WAND_DEF) |
1618 |
||
1619 |
lemma PBITBOP_WAND: "IN WAND PBITBOP" |
|
1620 |
by (import bword_bitop PBITBOP_WAND) |
|
1621 |
||
1622 |
consts |
|
1623 |
WOR :: "bool word => bool word => bool word" |
|
1624 |
||
| 17644 | 1625 |
specification (WOR) WOR_DEF: "ALL (l1::bool list) l2::bool list. |
1626 |
WOR (WORD l1) (WORD l2) = WORD (map2 op | l1 l2)" |
|
| 14516 | 1627 |
by (import bword_bitop WOR_DEF) |
1628 |
||
1629 |
lemma PBITBOP_WOR: "IN WOR PBITBOP" |
|
1630 |
by (import bword_bitop PBITBOP_WOR) |
|
1631 |
||
1632 |
consts |
|
1633 |
WXOR :: "bool word => bool word => bool word" |
|
1634 |
||
| 17644 | 1635 |
specification (WXOR) WXOR_DEF: "ALL (l1::bool list) l2::bool list. |
1636 |
WXOR (WORD l1) (WORD l2) = WORD (map2 (%(x::bool) y::bool. x ~= y) l1 l2)" |
|
| 14516 | 1637 |
by (import bword_bitop WXOR_DEF) |
1638 |
||
1639 |
lemma PBITBOP_WXOR: "IN WXOR PBITBOP" |
|
1640 |
by (import bword_bitop PBITBOP_WXOR) |
|
1641 |
||
1642 |
;end_setup |
|
1643 |
||
1644 |
end |
|
1645 |