diff -r 5e64a54f6790 -r 438d583ab378 src/Doc/System/Sessions.thy --- a/src/Doc/System/Sessions.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/Doc/System/Sessions.thy Sun May 12 14:41:13 2024 +0200 @@ -516,15 +516,24 @@ \<^medskip> Option \<^verbatim>\-H\ augments the cluster hosts to be used in this build process. - Each \<^verbatim>\-H\ option accepts multiple host names (separated by commas), which - all share the same (optional) parameters. Multiple \<^verbatim>\-H\ options may be - given to specify further hosts (with different parameters). For example: - \<^verbatim>\-H host1,host2:jobs=2,threads=4 -H host3:jobs=4,threads=6\. + Each \<^verbatim>\-H\ option accepts multiple host or cluster names (separated by + commas), which all share the same (optional) parameters or system options. + Multiple \<^verbatim>\-H\ options may be given to specify further hosts (with different + parameters). For example: \<^verbatim>\-H host1,host2:jobs=2,threads=4 -H + host3:jobs=4,threads=6\. The syntax for host parameters follows Isabelle outer syntax, notably with double-quoted string literals. On the command-line, this may require extra single quotes or escapes. For example: \<^verbatim>\-H 'host4:dirs="..."'\. + The system registry (\secref{sec:system-registry}) may define host and + cluster names in its tables \<^verbatim>\host\ and \<^verbatim>\cluster\, respectively. A name in + option \<^verbatim>\-H\ without prefix refers to the registry table \<^verbatim>\host\: each entry + consists of an optional \<^verbatim>\hostname\ and further options. A name with the + prefix ``\<^verbatim>\cluster.\'' refers to the table \<^verbatim>\cluster\: each entry provides + \<^verbatim>\hosts\, an array of names for entries of the table \<^verbatim>\host\ as above, and + additional options that apply to all hosts simultaneously. + The local host only participates in cluster build, if an explicit option \<^verbatim>\-j\ > 0 is given. The default is 0, which means that \<^verbatim>\isabelle build -H\ will initialize the build queue and oversee remote workers, but not run any @@ -602,6 +611,21 @@ @{verbatim [display] \ isabelle build -a -j2 -o threads=8 \ -H host1:jobs=2,threads=8 -H host2:jobs=4:threads=4,numa,shared\} + + \<^smallskip> + Use build hosts and cluster specifications: + @{verbatim [display] \ isabelle build -H a -H b -H cluster.xy\} + + The above requires a \<^path>\$ISABELLE_HOME_USER/etc/registry.toml\ file like + this: + +@{verbatim [display] \ host.a = { hostname = "host-a.acme.org", jobs = 2 } + host.b = { hostname = "host-b.acme.org", jobs = 2 } + + host.x = { hostname = "server1.example.com" } + host.y = { hostname = "server2.example.com" } + cluster.xy = { hosts = ["x", "y"], jobs = 4 } +\} \