author | wenzelm |
Mon, 07 May 2018 19:40:55 +0200 | |
changeset 68106 | a514e29db980 |
parent 68094 | 0b66aca9c965 |
child 68108 | 2277fe496d78 |
permissions | -rw-r--r-- |
54439 | 1 |
/* Title: Pure/General/bytes.scala |
2 |
Author: Makarius |
|
3 |
||
4 |
Immutable byte vectors versus UTF8 strings. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
64004 | 10 |
import java.io.{File => JFile, ByteArrayOutputStream, ByteArrayInputStream, |
64229 | 11 |
OutputStream, InputStream, FileInputStream, FileOutputStream} |
65070 | 12 |
import java.net.URL |
68094 | 13 |
import java.util.Base64 |
64004 | 14 |
|
15 |
import org.tukaani.xz.{XZInputStream, XZOutputStream} |
|
54440 | 16 |
|
17 |
||
54439 | 18 |
object Bytes |
19 |
{ |
|
20 |
val empty: Bytes = new Bytes(Array[Byte](), 0, 0) |
|
21 |
||
22 |
def apply(s: CharSequence): Bytes = |
|
23 |
{ |
|
24 |
val str = s.toString |
|
25 |
if (str.isEmpty) empty |
|
26 |
else { |
|
62527 | 27 |
val b = UTF8.bytes(str) |
54442
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
28 |
new Bytes(b, 0, b.length) |
54439 | 29 |
} |
30 |
} |
|
54440 | 31 |
|
63779 | 32 |
def apply(a: Array[Byte]): Bytes = apply(a, 0, a.length) |
33 |
||
54442
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
34 |
def apply(a: Array[Byte], offset: Int, length: Int): Bytes = |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
35 |
if (length == 0) empty |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
36 |
else { |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
37 |
val b = new Array[Byte](length) |
55618 | 38 |
System.arraycopy(a, offset, b, 0, length) |
54442
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
39 |
new Bytes(b, 0, b.length) |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
40 |
} |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
41 |
|
54440 | 42 |
|
68087 | 43 |
def hex(s: String): Bytes = |
44 |
{ |
|
45 |
def err(): Nothing = error("Malformed hexadecimal representation of bytes\n" + s) |
|
46 |
val len = s.length |
|
47 |
if (len % 2 != 0) err() |
|
48 |
||
49 |
val n = len / 2 |
|
50 |
val a = new Array[Byte](n) |
|
51 |
for (i <- 0 until n) { |
|
52 |
val j = 2 * i |
|
53 |
try { a(i) = Integer.parseInt(s.substring(j, j + 2), 16).toByte } |
|
54 |
catch { case _: NumberFormatException => err() } |
|
55 |
} |
|
56 |
new Bytes(a, 0, n) |
|
57 |
} |
|
58 |
||
59 |
||
54440 | 60 |
/* read */ |
61 |
||
64005
f6e965cf1617
clarified magic values (see also java/io/BufferedInputStream.java);
wenzelm
parents:
64004
diff
changeset
|
62 |
def read_stream(stream: InputStream, limit: Int = Integer.MAX_VALUE, hint: Int = 1024): Bytes = |
64004 | 63 |
if (limit == 0) empty |
64 |
else { |
|
64005
f6e965cf1617
clarified magic values (see also java/io/BufferedInputStream.java);
wenzelm
parents:
64004
diff
changeset
|
65 |
val out = new ByteArrayOutputStream(if (limit == Integer.MAX_VALUE) hint else limit) |
f6e965cf1617
clarified magic values (see also java/io/BufferedInputStream.java);
wenzelm
parents:
64004
diff
changeset
|
66 |
val buf = new Array[Byte](8192) |
64004 | 67 |
var m = 0 |
54440 | 68 |
|
69 |
do { |
|
64004 | 70 |
m = stream.read(buf, 0, buf.size min (limit - out.size)) |
71 |
if (m != -1) out.write(buf, 0, m) |
|
72 |
} while (m != -1 && limit > out.size) |
|
73 |
||
74 |
new Bytes(out.toByteArray, 0, out.size) |
|
54440 | 75 |
} |
64001
7ecb22be8f03
more general read_stream: return actual byte count;
wenzelm
parents:
63779
diff
changeset
|
76 |
|
67805
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
77 |
def read_block(stream: InputStream, length: Int): Option[Bytes] = |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
78 |
{ |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
79 |
val bytes = read_stream(stream, limit = length) |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
80 |
if (bytes.length == length) Some(bytes) else None |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
81 |
} |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
82 |
|
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
83 |
def read_line(stream: InputStream): Option[Bytes] = |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
84 |
{ |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
85 |
val out = new ByteArrayOutputStream(100) |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
86 |
var c = 0 |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
87 |
while ({ c = stream.read; c != -1 && c != 10 }) out.write(c) |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
88 |
|
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
89 |
if (c == -1 && out.size == 0) None |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
90 |
else { |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
91 |
val a = out.toByteArray |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
92 |
val n = a.length |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
93 |
val b = if (n > 0 && a(n - 1) == 13) a.take(n - 1) else a |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
94 |
Some(new Bytes(b, 0, b.length)) |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
95 |
} |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
96 |
} |
2d9a265b294e
more uniform Bytes.read_line/read_block operations;
wenzelm
parents:
65630
diff
changeset
|
97 |
|
64001
7ecb22be8f03
more general read_stream: return actual byte count;
wenzelm
parents:
63779
diff
changeset
|
98 |
def read(file: JFile): Bytes = |
64004 | 99 |
using(new FileInputStream(file))(read_stream(_, file.length.toInt)) |
64229 | 100 |
|
101 |
def read(path: Path): Bytes = read(path.file) |
|
102 |
||
65070 | 103 |
def read(url: URL): Bytes = using(url.openStream)(read_stream(_)) |
104 |
||
64229 | 105 |
|
106 |
/* write */ |
|
107 |
||
108 |
def write(file: JFile, bytes: Bytes) |
|
109 |
{ |
|
110 |
val stream = new FileOutputStream(file) |
|
111 |
try { bytes.write_stream(stream) } finally { stream.close } |
|
112 |
} |
|
113 |
||
114 |
def write(path: Path, bytes: Bytes): Unit = write(path.file, bytes) |
|
54439 | 115 |
} |
116 |
||
117 |
final class Bytes private( |
|
118 |
protected val bytes: Array[Byte], |
|
119 |
protected val offset: Int, |
|
60833 | 120 |
val length: Int) extends CharSequence |
54439 | 121 |
{ |
122 |
/* equality */ |
|
123 |
||
54440 | 124 |
override def equals(that: Any): Boolean = |
125 |
{ |
|
126 |
that match { |
|
127 |
case other: Bytes => |
|
128 |
if (this eq other) true |
|
129 |
else if (length != other.length) false |
|
130 |
else (0 until length).forall(i => bytes(offset + i) == other.bytes(other.offset + i)) |
|
131 |
case _ => false |
|
132 |
} |
|
133 |
} |
|
134 |
||
54439 | 135 |
private lazy val hash: Int = |
136 |
{ |
|
137 |
var h = 0 |
|
138 |
for (i <- offset until offset + length) { |
|
139 |
val b = bytes(i).asInstanceOf[Int] & 0xFF |
|
140 |
h = 31 * h + b |
|
141 |
} |
|
142 |
h |
|
143 |
} |
|
144 |
||
145 |
override def hashCode(): Int = hash |
|
146 |
||
147 |
||
148 |
/* content */ |
|
149 |
||
54512 | 150 |
lazy val sha1_digest: SHA1.Digest = SHA1.digest(bytes) |
54440 | 151 |
|
65279
fa62e095d8f1
clarified signature (again, see also 3ed43cfc8b14);
wenzelm
parents:
65070
diff
changeset
|
152 |
def text: String = |
fa62e095d8f1
clarified signature (again, see also 3ed43cfc8b14);
wenzelm
parents:
65070
diff
changeset
|
153 |
UTF8.decode_chars(s => s, bytes, offset, offset + length).toString |
fa62e095d8f1
clarified signature (again, see also 3ed43cfc8b14);
wenzelm
parents:
65070
diff
changeset
|
154 |
|
68094 | 155 |
def base64: String = |
156 |
{ |
|
157 |
val b = |
|
158 |
if (offset == 0 && length == bytes.length) bytes |
|
159 |
else Bytes(bytes, offset, length).bytes |
|
160 |
Base64.getEncoder.encodeToString(b) |
|
161 |
} |
|
162 |
||
68106 | 163 |
def maybe_base64: (Boolean, String) = |
164 |
{ |
|
165 |
val s = text |
|
166 |
if (this == Bytes(s)) (false, s) else (true, base64) |
|
167 |
} |
|
168 |
||
54444 | 169 |
override def toString: String = |
64224 | 170 |
{ |
65279
fa62e095d8f1
clarified signature (again, see also 3ed43cfc8b14);
wenzelm
parents:
65070
diff
changeset
|
171 |
val str = text |
64224 | 172 |
if (str.contains('\uFFFD')) "Bytes(" + length + ")" else str |
173 |
} |
|
54439 | 174 |
|
54440 | 175 |
def isEmpty: Boolean = length == 0 |
54439 | 176 |
|
65630 | 177 |
def proper: Option[Bytes] = if (isEmpty) None else Some(this) |
178 |
def proper_text: Option[String] = if (isEmpty) None else Some(text) |
|
179 |
||
54439 | 180 |
def +(other: Bytes): Bytes = |
54440 | 181 |
if (other.isEmpty) this |
182 |
else if (isEmpty) other |
|
54439 | 183 |
else { |
184 |
val new_bytes = new Array[Byte](length + other.length) |
|
55618 | 185 |
System.arraycopy(bytes, offset, new_bytes, 0, length) |
186 |
System.arraycopy(other.bytes, other.offset, new_bytes, length, other.length) |
|
54439 | 187 |
new Bytes(new_bytes, 0, new_bytes.length) |
188 |
} |
|
54440 | 189 |
|
190 |
||
60833 | 191 |
/* CharSequence operations */ |
192 |
||
193 |
def charAt(i: Int): Char = |
|
194 |
if (0 <= i && i < length) (bytes(offset + i).asInstanceOf[Int] & 0xFF).asInstanceOf[Char] |
|
195 |
else throw new IndexOutOfBoundsException |
|
196 |
||
197 |
def subSequence(i: Int, j: Int): Bytes = |
|
198 |
{ |
|
199 |
if (0 <= i && i <= j && j <= length) new Bytes(bytes, offset + i, j - i) |
|
200 |
else throw new IndexOutOfBoundsException |
|
201 |
} |
|
202 |
||
203 |
||
64004 | 204 |
/* streams */ |
205 |
||
206 |
def stream(): ByteArrayInputStream = new ByteArrayInputStream(bytes, offset, length) |
|
207 |
||
208 |
def write_stream(stream: OutputStream): Unit = stream.write(bytes, offset, length) |
|
209 |
||
210 |
||
211 |
/* XZ data compression */ |
|
54440 | 212 |
|
68018 | 213 |
def uncompress(cache: XZ.Cache = XZ.cache()): Bytes = |
214 |
using(new XZInputStream(stream(), cache))(Bytes.read_stream(_, hint = length)) |
|
64004 | 215 |
|
68018 | 216 |
def compress(options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.cache()): Bytes = |
64004 | 217 |
{ |
218 |
val result = new ByteArrayOutputStream(length) |
|
68018 | 219 |
using(new XZOutputStream(result, options, cache))(write_stream(_)) |
64004 | 220 |
new Bytes(result.toByteArray, 0, result.size) |
221 |
} |
|
54439 | 222 |
} |