| author | wenzelm | 
| Tue, 19 Dec 2023 19:18:09 +0100 | |
| changeset 79310 | dc6b58da806e | 
| parent 76183 | 8089593a364a | 
| child 80565 | 2acdaa0a7d29 | 
| permissions | -rw-r--r-- | 
| 76183 | 1  | 
(* Title: Pure/PIDE/byte_message.ML  | 
| 69448 | 2  | 
Author: Makarius  | 
3  | 
||
4  | 
Byte-oriented messages.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature BYTE_MESSAGE =  | 
|
8  | 
sig  | 
|
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
9  | 
val write: BinIO.outstream -> Bytes.T list -> unit  | 
| 73558 | 10  | 
val write_yxml: BinIO.outstream -> XML.tree -> unit  | 
| 69451 | 11  | 
val flush: BinIO.outstream -> unit  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
12  | 
val write_line: BinIO.outstream -> Bytes.T -> unit  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
13  | 
val read: BinIO.instream -> int -> Bytes.T  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
14  | 
val read_block: BinIO.instream -> int -> Bytes.T option * int  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
15  | 
val read_line: BinIO.instream -> Bytes.T option  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
16  | 
val write_message: BinIO.outstream -> Bytes.T list -> unit  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
17  | 
val write_message_string: BinIO.outstream -> string list -> unit  | 
| 73559 | 18  | 
val write_message_yxml: BinIO.outstream -> XML.body list -> unit  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
19  | 
val read_message: BinIO.instream -> Bytes.T list option  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
20  | 
val read_message_string: BinIO.instream -> string list option  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
21  | 
val write_line_message: BinIO.outstream -> Bytes.T -> unit  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
22  | 
val read_line_message: BinIO.instream -> Bytes.T option  | 
| 69448 | 23  | 
end;  | 
24  | 
||
25  | 
structure Byte_Message: BYTE_MESSAGE =  | 
|
26  | 
struct  | 
|
27  | 
||
| 69451 | 28  | 
(* output operations *)  | 
29  | 
||
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
30  | 
val write = List.app o Bytes.write_stream;  | 
| 69451 | 31  | 
|
| 75615 | 32  | 
fun write_yxml stream tree =  | 
33  | 
YXML.traverse (fn s => fn () => File_Stream.output stream s) tree ();  | 
|
| 73558 | 34  | 
|
| 69451 | 35  | 
fun flush stream = ignore (try BinIO.flushOut stream);  | 
36  | 
||
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
37  | 
fun write_line stream bs = (write stream [bs, Bytes.newline]; flush stream);  | 
| 69467 | 38  | 
|
| 69451 | 39  | 
|
40  | 
(* input operations *)  | 
|
41  | 
||
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
42  | 
fun read stream n = Bytes.read_stream n stream;  | 
| 69449 | 43  | 
|
44  | 
fun read_block stream n =  | 
|
| 69452 | 45  | 
let  | 
46  | 
val msg = read stream n;  | 
|
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
47  | 
val len = Bytes.length msg;  | 
| 69452 | 48  | 
in (if len = n then SOME msg else NONE, len) end;  | 
| 69449 | 49  | 
|
| 69448 | 50  | 
fun read_line stream =  | 
51  | 
let  | 
|
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
52  | 
val result = SOME o Bytes.trim_line;  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
53  | 
fun read_body bs =  | 
| 69448 | 54  | 
(case BinIO.input1 stream of  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
55  | 
NONE => if Bytes.is_empty bs then NONE else result bs  | 
| 69448 | 56  | 
| SOME b =>  | 
57  | 
(case Byte.byteToChar b of  | 
|
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
58  | 
#"\n" => result bs  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
59  | 
| c => read_body (Bytes.add (str c) bs)));  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
60  | 
in read_body Bytes.empty end;  | 
| 69448 | 61  | 
|
| 69449 | 62  | 
|
| 69454 | 63  | 
(* messages with multiple chunks (arbitrary content) *)  | 
| 69451 | 64  | 
|
| 69848 | 65  | 
fun make_header ns =  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
66  | 
[Bytes.string (space_implode "," (map Value.print_int ns)), Bytes.newline];  | 
| 69451 | 67  | 
|
| 69454 | 68  | 
fun write_message stream chunks =  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
69  | 
(write stream (make_header (map Bytes.length chunks) @ chunks); flush stream);  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
70  | 
|
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
71  | 
fun write_message_string stream =  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
72  | 
write_message stream o map Bytes.string;  | 
| 69451 | 73  | 
|
| 73559 | 74  | 
fun write_message_yxml stream chunks =  | 
75  | 
(write stream (make_header (map YXML.body_size chunks));  | 
|
76  | 
(List.app o List.app) (write_yxml stream) chunks;  | 
|
77  | 
flush stream);  | 
|
78  | 
||
| 69451 | 79  | 
fun parse_header line =  | 
80  | 
map Value.parse_nat (space_explode "," line)  | 
|
| 69454 | 81  | 
    handle Fail _ => error ("Malformed message header: " ^ quote line);
 | 
| 69451 | 82  | 
|
| 69449 | 83  | 
fun read_chunk stream n =  | 
| 69452 | 84  | 
(case read_block stream n of  | 
85  | 
(SOME chunk, _) => chunk  | 
|
86  | 
| (NONE, len) =>  | 
|
| 69449 | 87  | 
      error ("Malformed message chunk: unexpected EOF after " ^
 | 
| 69452 | 88  | 
string_of_int len ^ " of " ^ string_of_int n ^ " bytes"));  | 
| 69449 | 89  | 
|
90  | 
fun read_message stream =  | 
|
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
91  | 
read_line stream |> Option.map (Bytes.content #> parse_header #> map (read_chunk stream));  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
92  | 
|
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
93  | 
fun read_message_string stream =  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
94  | 
read_message stream |> (Option.map o map) Bytes.content;  | 
| 69451 | 95  | 
|
96  | 
||
97  | 
(* hybrid messages: line or length+block (with content restriction) *)  | 
|
98  | 
||
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
99  | 
(* line message format *)  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
100  | 
|
| 69452 | 101  | 
fun is_length msg =  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
102  | 
not (Bytes.is_empty msg) andalso Bytes.forall_string Symbol.is_ascii_digit msg;  | 
| 69451 | 103  | 
|
| 69452 | 104  | 
fun is_terminated msg =  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
105  | 
(case Bytes.last_string msg of  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
106  | 
NONE => false  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
107  | 
| SOME s => Symbol.is_ascii_line_terminator s);  | 
| 69451 | 108  | 
|
109  | 
fun write_line_message stream msg =  | 
|
| 69452 | 110  | 
if is_length msg orelse is_terminated msg then  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
111  | 
    error ("Bad content for line message:\n" ^ Bytes.beginning 100 msg)
 | 
| 69451 | 112  | 
else  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
113  | 
let val n = Bytes.length msg in  | 
| 69454 | 114  | 
write stream  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
115  | 
((if n > 100 orelse Bytes.exists_string (fn s => s = "\n") msg  | 
| 
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
116  | 
then make_header [n + 1] else []) @ [msg, Bytes.newline]);  | 
| 69451 | 117  | 
flush stream  | 
118  | 
end;  | 
|
119  | 
||
120  | 
fun read_line_message stream =  | 
|
121  | 
(case read_line stream of  | 
|
| 69452 | 122  | 
NONE => NONE  | 
123  | 
| SOME line =>  | 
|
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
124  | 
(case try (Value.parse_nat o Bytes.content) line of  | 
| 69451 | 125  | 
NONE => SOME line  | 
| 
75577
 
c51e1cef1eae
more scalable byte messages, notably for Scala functions in ML;
 
wenzelm 
parents: 
75571 
diff
changeset
 | 
126  | 
| SOME n => Option.map Bytes.trim_line (#1 (read_block stream n))));  | 
| 69448 | 127  | 
|
128  | 
end;  |