author | wenzelm |
Wed, 22 Jun 2022 13:42:30 +0200 | |
changeset 75588 | 3349e360b71d |
parent 75587 | 79b4efd17d2b |
child 75620 | 44815dc2b8f9 |
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 |
||
75393 | 18 |
object Bytes { |
54439 | 19 |
val empty: Bytes = new Bytes(Array[Byte](), 0, 0) |
20 |
||
75393 | 21 |
def apply(s: CharSequence): Bytes = { |
54439 | 22 |
val str = s.toString |
23 |
if (str.isEmpty) empty |
|
24 |
else { |
|
62527 | 25 |
val b = UTF8.bytes(str) |
54442
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
26 |
new Bytes(b, 0, b.length) |
54439 | 27 |
} |
28 |
} |
|
54440 | 29 |
|
63779 | 30 |
def apply(a: Array[Byte]): Bytes = apply(a, 0, a.length) |
31 |
||
54442
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
32 |
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
|
33 |
if (length == 0) empty |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
34 |
else { |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
35 |
val b = new Array[Byte](length) |
55618 | 36 |
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
|
37 |
new Bytes(b, 0, b.length) |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
38 |
} |
c39972ddd672
more specific Protocol_Output: empty message.body, main content via bytes/text;
wenzelm
parents:
54440
diff
changeset
|
39 |
|
69454 | 40 |
val newline: Bytes = apply("\n") |
54440 | 41 |
|
73576 | 42 |
|
43 |
/* base64 */ |
|
44 |
||
75587 | 45 |
def decode_base64(s: String): Bytes = { |
68108 | 46 |
val a = Base64.getDecoder.decode(s) |
47 |
new Bytes(a, 0, a.length) |
|
48 |
} |
|
49 |
||
75588 | 50 |
object Decode_Base64 extends Scala.Fun_Bytes("decode_base64") { |
73576 | 51 |
val here = Scala_Project.here |
75588 | 52 |
def apply(arg: Bytes): Bytes = decode_base64(arg.text) |
73576 | 53 |
} |
54 |
||
75588 | 55 |
object Encode_Base64 extends Scala.Fun_Bytes("encode_base64") { |
73576 | 56 |
val here = Scala_Project.here |
75588 | 57 |
def apply(arg: Bytes): Bytes = Bytes(arg.encode_base64) |
73576 | 58 |
} |
59 |
||
68087 | 60 |
|
75579 | 61 |
/* XZ compression */ |
62 |
||
75588 | 63 |
object Compress extends Scala.Fun_Bytes("compress") { |
75579 | 64 |
val here = Scala_Project.here |
75588 | 65 |
def apply(arg: Bytes): Bytes = arg.compress() |
75579 | 66 |
} |
67 |
||
75588 | 68 |
object Uncompress extends Scala.Fun_Bytes("uncompress") { |
75579 | 69 |
val here = Scala_Project.here |
75588 | 70 |
def apply(arg: Bytes): Bytes = arg.uncompress() |
75579 | 71 |
} |
72 |
||
73 |
||
54440 | 74 |
/* read */ |
75 |
||
64005
f6e965cf1617
clarified magic values (see also java/io/BufferedInputStream.java);
wenzelm
parents:
64004
diff
changeset
|
76 |
def read_stream(stream: InputStream, limit: Int = Integer.MAX_VALUE, hint: Int = 1024): Bytes = |
64004 | 77 |
if (limit == 0) empty |
78 |
else { |
|
73414 | 79 |
val out_size = (if (limit == Integer.MAX_VALUE) hint else limit) max 1024 |
80 |
val out = new ByteArrayOutputStream(out_size) |
|
64005
f6e965cf1617
clarified magic values (see also java/io/BufferedInputStream.java);
wenzelm
parents:
64004
diff
changeset
|
81 |
val buf = new Array[Byte](8192) |
64004 | 82 |
var m = 0 |
54440 | 83 |
|
75382
81673c441ce3
tuned: eliminted do-while for the sake of scala3;
wenzelm
parents:
73576
diff
changeset
|
84 |
var cont = true |
81673c441ce3
tuned: eliminted do-while for the sake of scala3;
wenzelm
parents:
73576
diff
changeset
|
85 |
while (cont) { |
73554 | 86 |
m = stream.read(buf, 0, buf.length min (limit - out.size)) |
64004 | 87 |
if (m != -1) out.write(buf, 0, m) |
75382
81673c441ce3
tuned: eliminted do-while for the sake of scala3;
wenzelm
parents:
73576
diff
changeset
|
88 |
cont = (m != -1 && limit > out.size) |
81673c441ce3
tuned: eliminted do-while for the sake of scala3;
wenzelm
parents:
73576
diff
changeset
|
89 |
} |
64004 | 90 |
|
91 |
new Bytes(out.toByteArray, 0, out.size) |
|
54440 | 92 |
} |
64001
7ecb22be8f03
more general read_stream: return actual byte count;
wenzelm
parents:
63779
diff
changeset
|
93 |
|
75393 | 94 |
def read(file: JFile): Bytes = { |
71152
f2d848a596d1
more robust: file length can be invalid in odd situations;
wenzelm
parents:
71151
diff
changeset
|
95 |
val length = file.length |
f2d848a596d1
more robust: file length can be invalid in odd situations;
wenzelm
parents:
71151
diff
changeset
|
96 |
val limit = if (length < 0 || length > Integer.MAX_VALUE) Integer.MAX_VALUE else length.toInt |
f2d848a596d1
more robust: file length can be invalid in odd situations;
wenzelm
parents:
71151
diff
changeset
|
97 |
using(new FileInputStream(file))(read_stream(_, limit = limit)) |
f2d848a596d1
more robust: file length can be invalid in odd situations;
wenzelm
parents:
71151
diff
changeset
|
98 |
} |
64229 | 99 |
|
100 |
def read(path: Path): Bytes = read(path.file) |
|
101 |
||
65070 | 102 |
def read(url: URL): Bytes = using(url.openStream)(read_stream(_)) |
103 |
||
64229 | 104 |
|
105 |
/* write */ |
|
106 |
||
69393
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
wenzelm
parents:
69365
diff
changeset
|
107 |
def write(file: JFile, bytes: Bytes): Unit = |
ed0824ef337e
static type for Library.using: avoid Java 11 warnings on "Illegal reflective access";
wenzelm
parents:
69365
diff
changeset
|
108 |
using(new FileOutputStream(file))(bytes.write_stream(_)) |
64229 | 109 |
|
110 |
def write(path: Path, bytes: Bytes): Unit = write(path.file, bytes) |
|
54439 | 111 |
} |
112 |
||
113 |
final class Bytes private( |
|
114 |
protected val bytes: Array[Byte], |
|
115 |
protected val offset: Int, |
|
75393 | 116 |
val length: Int) extends CharSequence { |
54439 | 117 |
/* equality */ |
118 |
||
75393 | 119 |
override def equals(that: Any): Boolean = { |
54440 | 120 |
that match { |
121 |
case other: Bytes => |
|
122 |
if (this eq other) true |
|
123 |
else if (length != other.length) false |
|
124 |
else (0 until length).forall(i => bytes(offset + i) == other.bytes(other.offset + i)) |
|
125 |
case _ => false |
|
126 |
} |
|
127 |
} |
|
128 |
||
75393 | 129 |
private lazy val hash: Int = { |
54439 | 130 |
var h = 0 |
131 |
for (i <- offset until offset + length) { |
|
132 |
val b = bytes(i).asInstanceOf[Int] & 0xFF |
|
133 |
h = 31 * h + b |
|
134 |
} |
|
135 |
h |
|
136 |
} |
|
137 |
||
138 |
override def hashCode(): Int = hash |
|
139 |
||
140 |
||
141 |
/* content */ |
|
142 |
||
54512 | 143 |
lazy val sha1_digest: SHA1.Digest = SHA1.digest(bytes) |
54440 | 144 |
|
69448 | 145 |
def is_empty: Boolean = length == 0 |
146 |
||
147 |
def iterator: Iterator[Byte] = |
|
148 |
for (i <- (offset until (offset + length)).iterator) |
|
149 |
yield bytes(i) |
|
150 |
||
75393 | 151 |
def array: Array[Byte] = { |
69365
c5b3860d29ef
avoid loading of font file, to eliminate "Illegal reflective access by com.lowagie.text.pdf.MappedRandomAccessFile$1 (iText-2.1.5.jar) to method java.nio.DirectByteBuffer.cleaner()" -- due to com.lowagie.text.pdf.TrueTypeFont.process() / RandomAccessFileOrArray;
wenzelm
parents:
68167
diff
changeset
|
152 |
val a = new Array[Byte](length) |
c5b3860d29ef
avoid loading of font file, to eliminate "Illegal reflective access by com.lowagie.text.pdf.MappedRandomAccessFile$1 (iText-2.1.5.jar) to method java.nio.DirectByteBuffer.cleaner()" -- due to com.lowagie.text.pdf.TrueTypeFont.process() / RandomAccessFileOrArray;
wenzelm
parents:
68167
diff
changeset
|
153 |
System.arraycopy(bytes, offset, a, 0, length) |
c5b3860d29ef
avoid loading of font file, to eliminate "Illegal reflective access by com.lowagie.text.pdf.MappedRandomAccessFile$1 (iText-2.1.5.jar) to method java.nio.DirectByteBuffer.cleaner()" -- due to com.lowagie.text.pdf.TrueTypeFont.process() / RandomAccessFileOrArray;
wenzelm
parents:
68167
diff
changeset
|
154 |
a |
c5b3860d29ef
avoid loading of font file, to eliminate "Illegal reflective access by com.lowagie.text.pdf.MappedRandomAccessFile$1 (iText-2.1.5.jar) to method java.nio.DirectByteBuffer.cleaner()" -- due to com.lowagie.text.pdf.TrueTypeFont.process() / RandomAccessFileOrArray;
wenzelm
parents:
68167
diff
changeset
|
155 |
} |
c5b3860d29ef
avoid loading of font file, to eliminate "Illegal reflective access by com.lowagie.text.pdf.MappedRandomAccessFile$1 (iText-2.1.5.jar) to method java.nio.DirectByteBuffer.cleaner()" -- due to com.lowagie.text.pdf.TrueTypeFont.process() / RandomAccessFileOrArray;
wenzelm
parents:
68167
diff
changeset
|
156 |
|
73561
c83152933579
clarified signature: Bytes extends CharSequence already (see d201996f72a8);
wenzelm
parents:
73559
diff
changeset
|
157 |
def text: String = UTF8.decode_permissive(this) |
65279
fa62e095d8f1
clarified signature (again, see also 3ed43cfc8b14);
wenzelm
parents:
65070
diff
changeset
|
158 |
|
75587 | 159 |
def encode_base64: String = { |
68094 | 160 |
val b = |
161 |
if (offset == 0 && length == bytes.length) bytes |
|
162 |
else Bytes(bytes, offset, length).bytes |
|
163 |
Base64.getEncoder.encodeToString(b) |
|
164 |
} |
|
165 |
||
75587 | 166 |
def maybe_encode_base64: (Boolean, String) = { |
68106 | 167 |
val s = text |
75587 | 168 |
if (this == Bytes(s)) (false, s) else (true, encode_base64) |
68106 | 169 |
} |
170 |
||
68150
f0f34cbed539
clarified output: avoid costly operations on huge blobs;
wenzelm
parents:
68149
diff
changeset
|
171 |
override def toString: String = "Bytes(" + length + ")" |
54439 | 172 |
|
72885 | 173 |
def proper: Option[Bytes] = if (is_empty) None else Some(this) |
174 |
def proper_text: Option[String] = if (is_empty) None else Some(text) |
|
65630 | 175 |
|
54439 | 176 |
def +(other: Bytes): Bytes = |
72885 | 177 |
if (other.is_empty) this |
178 |
else if (is_empty) other |
|
54439 | 179 |
else { |
180 |
val new_bytes = new Array[Byte](length + other.length) |
|
55618 | 181 |
System.arraycopy(bytes, offset, new_bytes, 0, length) |
182 |
System.arraycopy(other.bytes, other.offset, new_bytes, length, other.length) |
|
54439 | 183 |
new Bytes(new_bytes, 0, new_bytes.length) |
184 |
} |
|
54440 | 185 |
|
186 |
||
60833 | 187 |
/* CharSequence operations */ |
188 |
||
189 |
def charAt(i: Int): Char = |
|
190 |
if (0 <= i && i < length) (bytes(offset + i).asInstanceOf[Int] & 0xFF).asInstanceOf[Char] |
|
191 |
else throw new IndexOutOfBoundsException |
|
192 |
||
75393 | 193 |
def subSequence(i: Int, j: Int): Bytes = { |
60833 | 194 |
if (0 <= i && i <= j && j <= length) new Bytes(bytes, offset + i, j - i) |
195 |
else throw new IndexOutOfBoundsException |
|
196 |
} |
|
197 |
||
69448 | 198 |
def trim_line: Bytes = |
199 |
if (length >= 2 && charAt(length - 2) == 13 && charAt(length - 1) == 10) |
|
200 |
subSequence(0, length - 2) |
|
201 |
else if (length >= 1 && (charAt(length - 1) == 13 || charAt(length - 1) == 10)) |
|
202 |
subSequence(0, length - 1) |
|
203 |
else this |
|
204 |
||
60833 | 205 |
|
64004 | 206 |
/* streams */ |
207 |
||
208 |
def stream(): ByteArrayInputStream = new ByteArrayInputStream(bytes, offset, length) |
|
209 |
||
210 |
def write_stream(stream: OutputStream): Unit = stream.write(bytes, offset, length) |
|
211 |
||
212 |
||
213 |
/* XZ data compression */ |
|
54440 | 214 |
|
73024 | 215 |
def uncompress(cache: XZ.Cache = XZ.Cache()): Bytes = |
68018 | 216 |
using(new XZInputStream(stream(), cache))(Bytes.read_stream(_, hint = length)) |
64004 | 217 |
|
75393 | 218 |
def compress(options: XZ.Options = XZ.options(), cache: XZ.Cache = XZ.Cache()): Bytes = { |
64004 | 219 |
val result = new ByteArrayOutputStream(length) |
68018 | 220 |
using(new XZOutputStream(result, options, cache))(write_stream(_)) |
64004 | 221 |
new Bytes(result.toByteArray, 0, result.size) |
222 |
} |
|
68167 | 223 |
|
75393 | 224 |
def maybe_compress( |
225 |
options: XZ.Options = XZ.options(), |
|
226 |
cache: XZ.Cache = XZ.Cache() |
|
227 |
) : (Boolean, Bytes) = { |
|
68167 | 228 |
val compressed = compress(options = options, cache = cache) |
229 |
if (compressed.length < length) (true, compressed) else (false, this) |
|
230 |
} |
|
54439 | 231 |
} |