| author | wenzelm | 
| Tue, 09 Jan 2018 15:18:41 +0100 | |
| changeset 67385 | deb9b0283259 | 
| parent 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 | 
| 67120 | 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]: | 
| 67120 | 13 | fixes x y :: int | 
| 54427 
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) | 
| 67120 | 18 | case 1 | 
| 19 | then show ?case by simp | |
| 20 | next | |
| 21 | case 2 | |
| 22 | then show ?case by (simp only: Min_def) | |
| 23 | next | |
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 24 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 25 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 26 | proof (cases y rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 27 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 28 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 29 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 30 | then have "0 \<le> bin AND bin'" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 31 | with 1 show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 32 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 33 | qed | 
| 67120 | 34 | qed | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 35 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 36 | lemma OR_lower [simp]: | 
| 67120 | 37 | fixes x y :: int | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 38 | assumes "0 \<le> x" "0 \<le> y" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 39 | shows "0 \<le> x OR y" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 40 | using assms | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 41 | proof (induct x arbitrary: y rule: bin_induct) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 42 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 43 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 44 | proof (cases y rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 45 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 46 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 47 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 48 | moreover from 1 3 have "0 \<le> bin'" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 49 | by (cases bit') (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 50 | ultimately have "0 \<le> bin OR bin'" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 51 | with 1 show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 52 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 53 | qed | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 54 | qed simp_all | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 55 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 56 | lemma XOR_lower [simp]: | 
| 67120 | 57 | fixes x y :: int | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 58 | assumes "0 \<le> x" "0 \<le> y" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 59 | shows "0 \<le> x XOR y" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 60 | using assms | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 61 | proof (induct x arbitrary: y rule: bin_induct) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 62 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 63 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 64 | proof (cases y rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 65 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 66 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 67 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 68 | moreover from 1 3 have "0 \<le> bin'" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 69 | by (cases bit') (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 70 | ultimately have "0 \<le> bin XOR bin'" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 71 | with 1 show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 72 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 73 | qed | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 74 | next | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 75 | case 2 | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 76 | then show ?case by (simp only: Min_def) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 77 | qed simp | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 78 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 79 | lemma AND_upper1 [simp]: | 
| 67120 | 80 | fixes x y :: int | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 81 | assumes "0 \<le> x" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 82 | shows "x AND y \<le> x" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 83 | using assms | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 84 | proof (induct x arbitrary: y rule: bin_induct) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 85 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 86 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 87 | proof (cases y rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 88 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 89 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 90 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 91 | then have "bin AND bin' \<le> bin" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 92 | with 1 show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 93 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 94 | qed | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 95 | next | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 96 | case 2 | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 97 | then show ?case by (simp only: Min_def) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 98 | qed simp | 
| 
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 | lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 101 | lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 102 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 103 | lemma AND_upper2 [simp]: | 
| 67120 | 104 | fixes x y :: int | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 105 | assumes "0 \<le> y" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 106 | shows "x AND y \<le> y" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 107 | using assms | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 108 | proof (induct y arbitrary: x rule: bin_induct) | 
| 67120 | 109 | case 1 | 
| 110 | then show ?case by simp | |
| 111 | next | |
| 112 | case 2 | |
| 113 | then show ?case by (simp only: Min_def) | |
| 114 | next | |
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 115 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 116 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 117 | proof (cases x rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 118 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 119 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 120 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 121 | then have "bin' AND bin \<le> bin" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 122 | with 1 show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 123 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 124 | qed | 
| 67120 | 125 | qed | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 126 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 127 | lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 128 | lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 129 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 130 | lemma OR_upper: | 
| 67120 | 131 | fixes x y :: int | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 132 | assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 133 | shows "x OR y < 2 ^ n" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 134 | using assms | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 135 | proof (induct x arbitrary: y n rule: bin_induct) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 136 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 137 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 138 | proof (cases y rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 139 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 140 | show ?thesis | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 141 | proof (cases n) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 142 | case 0 | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 143 | with 3 have "bin BIT bit = 0" by simp | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 144 | then have "bin = 0" and "\<not> bit" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 145 | by (auto simp add: Bit_def split: if_splits) arith | 
| 61799 | 146 | 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 | 147 | by simp | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 148 | next | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 149 | case (Suc m) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 150 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 151 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 152 | moreover from 3 Suc have "bin < 2 ^ m" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 153 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 154 | moreover from 1 3 Suc have "bin' < 2 ^ m" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 155 | by (cases bit') (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 156 | ultimately have "bin OR bin' < 2 ^ m" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 157 | with 1 Suc show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 158 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 159 | qed | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 160 | qed | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 161 | qed simp_all | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 162 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 163 | lemma XOR_upper: | 
| 67120 | 164 | fixes x y :: int | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 165 | assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 166 | shows "x XOR y < 2 ^ n" | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 167 | using assms | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 168 | proof (induct x arbitrary: y n rule: bin_induct) | 
| 67120 | 169 | case 1 | 
| 170 | then show ?case by simp | |
| 171 | next | |
| 172 | case 2 | |
| 173 | then show ?case by (simp only: Min_def) | |
| 174 | next | |
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 175 | case (3 bin bit) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 176 | show ?case | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 177 | proof (cases y rule: bin_exhaust) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 178 | case (1 bin' bit') | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 179 | show ?thesis | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 180 | proof (cases n) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 181 | case 0 | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 182 | with 3 have "bin BIT bit = 0" by simp | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 183 | then have "bin = 0" and "\<not> bit" | 
| 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 184 | by (auto simp add: Bit_def split: if_splits) arith | 
| 61799 | 185 | 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 | 186 | by simp | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 187 | next | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 188 | case (Suc m) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 189 | from 3 have "0 \<le> bin" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 190 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 191 | moreover from 3 Suc have "bin < 2 ^ m" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 192 | by (cases bit) (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 193 | moreover from 1 3 Suc have "bin' < 2 ^ m" | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 194 | by (cases bit') (simp_all add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 195 | ultimately have "bin XOR bin' < 2 ^ m" by (rule 3) | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 196 | with 1 Suc show ?thesis | 
| 54847 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 haftmann parents: 
54427diff
changeset | 197 | by simp (simp add: Bit_def) | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 198 | qed | 
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 199 | qed | 
| 67120 | 200 | qed | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 201 | |
| 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: diff
changeset | 202 | end |