author | wenzelm |
Mon, 08 May 2017 10:27:13 +0200 | |
changeset 65767 | 222ed8901008 |
parent 61799 | 4cf66f21b764 |
child 67120 | 491fd7f0b5df |
permissions | -rw-r--r-- |
55210 | 1 |
(* Title: HOL/Word/Bit_Comparison.thy |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
2 |
Author: Stefan Berghofer |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
3 |
Copyright: secunet Security Networks AG |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
4 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
5 |
Comparison on bit operations on integers. |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
6 |
*) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
7 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
8 |
theory Bit_Comparison |
54854
3324a0078636
prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
haftmann
parents:
54847
diff
changeset
|
9 |
imports Bits_Int |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
10 |
begin |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
11 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
12 |
lemma AND_lower [simp]: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
13 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
14 |
assumes "0 \<le> x" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
15 |
shows "0 \<le> x AND y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
16 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
17 |
proof (induct x arbitrary: y rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
18 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
19 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
20 |
proof (cases y rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
21 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
22 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
23 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
24 |
then have "0 \<le> bin AND bin'" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
25 |
with 1 show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
26 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
27 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
28 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
29 |
case 2 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
30 |
then show ?case by (simp only: Min_def) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
31 |
qed simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
32 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
33 |
lemma OR_lower [simp]: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
34 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
35 |
assumes "0 \<le> x" "0 \<le> y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
36 |
shows "0 \<le> x OR y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
37 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
38 |
proof (induct x arbitrary: y rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
39 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
40 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
41 |
proof (cases y rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
42 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
43 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
44 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
45 |
moreover from 1 3 have "0 \<le> bin'" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
46 |
by (cases bit') (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
47 |
ultimately have "0 \<le> bin OR bin'" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
48 |
with 1 show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
49 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
50 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
51 |
qed simp_all |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
52 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
53 |
lemma XOR_lower [simp]: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
54 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
55 |
assumes "0 \<le> x" "0 \<le> y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
56 |
shows "0 \<le> x XOR y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
57 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
58 |
proof (induct x arbitrary: y rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
59 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
60 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
61 |
proof (cases y rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
62 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
63 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
64 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
65 |
moreover from 1 3 have "0 \<le> bin'" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
66 |
by (cases bit') (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
67 |
ultimately have "0 \<le> bin XOR bin'" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
68 |
with 1 show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
69 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
70 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
71 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
72 |
case 2 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
73 |
then show ?case by (simp only: Min_def) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
74 |
qed simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
75 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
76 |
lemma AND_upper1 [simp]: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
77 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
78 |
assumes "0 \<le> x" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
79 |
shows "x AND y \<le> x" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
80 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
81 |
proof (induct x arbitrary: y rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
82 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
83 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
84 |
proof (cases y rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
85 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
86 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
87 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
88 |
then have "bin AND bin' \<le> bin" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
89 |
with 1 show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
90 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
91 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
92 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
93 |
case 2 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
94 |
then show ?case by (simp only: Min_def) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
95 |
qed simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
96 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
97 |
lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
98 |
lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
99 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
100 |
lemma AND_upper2 [simp]: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
101 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
102 |
assumes "0 \<le> y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
103 |
shows "x AND y \<le> y" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
104 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
105 |
proof (induct y arbitrary: x rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
106 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
107 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
108 |
proof (cases x rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
109 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
110 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
111 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
112 |
then have "bin' AND bin \<le> bin" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
113 |
with 1 show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
114 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
115 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
116 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
117 |
case 2 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
118 |
then show ?case by (simp only: Min_def) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
119 |
qed simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
120 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
121 |
lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
122 |
lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
123 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
124 |
lemma OR_upper: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
125 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
126 |
assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
127 |
shows "x OR y < 2 ^ n" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
128 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
129 |
proof (induct x arbitrary: y n rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
130 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
131 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
132 |
proof (cases y rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
133 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
134 |
show ?thesis |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
135 |
proof (cases n) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
136 |
case 0 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
137 |
with 3 have "bin BIT bit = 0" by simp |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
138 |
then have "bin = 0" and "\<not> bit" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
139 |
by (auto simp add: Bit_def split: if_splits) arith |
61799 | 140 |
then show ?thesis using 0 1 \<open>y < 2 ^ n\<close> |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
141 |
by simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
142 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
143 |
case (Suc m) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
144 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
145 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
146 |
moreover from 3 Suc have "bin < 2 ^ m" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
147 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
148 |
moreover from 1 3 Suc have "bin' < 2 ^ m" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
149 |
by (cases bit') (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
150 |
ultimately have "bin OR bin' < 2 ^ m" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
151 |
with 1 Suc show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
152 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
153 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
154 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
155 |
qed simp_all |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
156 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
157 |
lemma XOR_upper: |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
158 |
fixes x :: int and y :: int |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
159 |
assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
160 |
shows "x XOR y < 2 ^ n" |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
161 |
using assms |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
162 |
proof (induct x arbitrary: y n rule: bin_induct) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
163 |
case (3 bin bit) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
164 |
show ?case |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
165 |
proof (cases y rule: bin_exhaust) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
166 |
case (1 bin' bit') |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
167 |
show ?thesis |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
168 |
proof (cases n) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
169 |
case 0 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
170 |
with 3 have "bin BIT bit = 0" by simp |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
171 |
then have "bin = 0" and "\<not> bit" |
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
172 |
by (auto simp add: Bit_def split: if_splits) arith |
61799 | 173 |
then show ?thesis using 0 1 \<open>y < 2 ^ n\<close> |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
174 |
by simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
175 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
176 |
case (Suc m) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
177 |
from 3 have "0 \<le> bin" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
178 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
179 |
moreover from 3 Suc have "bin < 2 ^ m" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
180 |
by (cases bit) (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
181 |
moreover from 1 3 Suc have "bin' < 2 ^ m" |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
182 |
by (cases bit') (simp_all add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
183 |
ultimately have "bin XOR bin' < 2 ^ m" by (rule 3) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
184 |
with 1 Suc show ?thesis |
54847
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
haftmann
parents:
54427
diff
changeset
|
185 |
by simp (simp add: Bit_def) |
54427
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
186 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
187 |
qed |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
188 |
next |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
189 |
case 2 |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
190 |
then show ?case by (simp only: Min_def) |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
191 |
qed simp |
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
192 |
|
783861a66a60
separated comparision on bit operations into separate theory
haftmann
parents:
diff
changeset
|
193 |
end |