src/Doc/System/Misc.thy
changeset 75555 197a5b3a1ea2
parent 75511 b32fdb67f851
child 75556 1f6fc2416a48
equal deleted inserted replaced
75554:be33ca6f45d7 75555:197a5b3a1ea2
   298   Both variants require SSH access to the target server, via public key
   298   Both variants require SSH access to the target server, via public key
   299   without password.
   299   without password.
   300 \<close>
   300 \<close>
   301 
   301 
   302 
   302 
   303 section \<open>Mercurial repository synchronization\<close>
   303 section \<open>Mercurial repository synchronization \label{sec:tool-hg-sync}\<close>
   304 
   304 
   305 text \<open>
   305 text \<open>
   306   The @{tool_def hg_sync} tool synchronizes a local Mercurial repository with
   306   The @{tool_def hg_sync} tool synchronizes a local Mercurial repository with
   307   a target directory.
   307   a target directory.
   308 
   308 
   423   derived Isabelle logos for their own projects using this template. The
   423   derived Isabelle logos for their own projects using this template. The
   424   license is the same as for the regular Isabelle distribution (BSD).
   424   license is the same as for the regular Isabelle distribution (BSD).
   425 \<close>
   425 \<close>
   426 
   426 
   427 
   427 
       
   428 section \<open>Isabelle and AFP repository synchronization\<close>
       
   429 
       
   430 text \<open>
       
   431   The @{tool_def sync} tool synchronizes a local Isabelle and AFP repository,
       
   432   potentially including its prebuilt \<^verbatim>\<open>.jar\<close> files and session images.
       
   433 
       
   434   @{verbatim [display]
       
   435 \<open>Usage: isabelle sync [OPTIONS] TARGET
       
   436 
       
   437   Options are:
       
   438     -A ROOT      include AFP with given root directory (":" for $AFP_BASE)
       
   439     -H           purge heaps directory on target
       
   440     -I NAME      include session heap image and build database
       
   441                  (based on accidental local state)
       
   442     -J           preserve *.jar files
       
   443     -S           robust (but less portable) treatment of spaces in
       
   444                  file and directory names on the target
       
   445     -T           thorough treatment of file content and directory times
       
   446     -a REV       explicit AFP revision (default: state of working directory)
       
   447     -n           no changes: dry-run
       
   448     -r REV       explicit revision (default: state of working directory)
       
   449     -p PORT      explicit SSH port (default: 22)
       
   450     -v           verbose
       
   451 
       
   452   Synchronize Isabelle + AFP repositories, based on "isabelle hg_sync".\<close>}
       
   453 
       
   454   The approach is to apply @{tool hg_sync} (\secref{sec:tool-hg-sync}) on the
       
   455   underlying Isabelle repository plus a given AFP repository (optional). This
       
   456   means that the Isabelle installation needs to be a Mercurial repository
       
   457   clone: a regular download of the Isabelle distribution is not sufficient!
       
   458 
       
   459   On the target side, AFP is placed into @{setting ISABELLE_HOME} as immediate
       
   460   sub-directory with the literal name \<^verbatim>\<open>AFP\<close>; thus it can be easily included
       
   461   elsewhere, e.g. @{tool build}~\<^verbatim>\<open>-d\<close>~\<^verbatim>\<open>'~~/AFP'\<close> on the remote side.
       
   462 
       
   463   \<^medskip> Options \<^verbatim>\<open>-T\<close>, \<^verbatim>\<open>-n\<close>, \<^verbatim>\<open>-p\<close>, \<^verbatim>\<open>-v\<close> are the same as the underlying
       
   464   @{tool hg_sync}.
       
   465 
       
   466   \<^medskip> Options \<^verbatim>\<open>-r\<close> and \<^verbatim>\<open>-a\<close> are the same as option \<^verbatim>\<open>-r\<close> for @{tool hg_sync},
       
   467   but for the Isabelle and AFP repositories, respectively. The AFP version is
       
   468   only used if a corresponding repository is given via option \<^verbatim>\<open>-A\<close>, either
       
   469   with explicit root directory, or as \<^verbatim>\<open>-A:\<close> to refer to \<^verbatim>\<open>$AFP_BASE\<close> (this
       
   470   assumes AFP as component of the local Isabelle installation).
       
   471 
       
   472   \<^medskip> Option \<^verbatim>\<open>-J\<close> uploads existing \<^verbatim>\<open>.jar\<close> files, which are usually
       
   473   Isabelle/Scala/Java modules under control of \<^verbatim>\<open>etc/build.props\<close> (see also
       
   474   \secref{sec:scala-build}). Normally, the underlying dependency management is
       
   475   accurate: bad uploads will be rebuilt on the remote side (or ignored). For
       
   476   historic Isabelle versions, going far back into the past via option \<^verbatim>\<open>-r\<close>,
       
   477   it is better to omit option \<^verbatim>\<open>-J\<close> and thus purge \<^verbatim>\<open>.jar\<close> files on the target
       
   478   (because they do not belong to the repository).
       
   479 
       
   480   \<^medskip> Option \<^verbatim>\<open>-I\<close> uploads a collection of session images. The set of \<^verbatim>\<open>-I\<close>
       
   481   options specifies the end-points in the session build graph, including all
       
   482   required ancestors. The result collection is uploaded using the underlying
       
   483   \<^verbatim>\<open>rsync\<close> policies: unchanged images are ignored. Session images are
       
   484   assembled within the target \<^verbatim>\<open>heaps\<close> directory: this scheme fits together
       
   485   with @{tool build}~\<^verbatim>\<open>-o system_heaps\<close>. Images are taken as-is from the local
       
   486   Isabelle installation, regardless of option \<^verbatim>\<open>-r\<close>. Bad images do not cause
       
   487   any harm, apart from wasting network bandwidth and file-system space:
       
   488   running e.g. @{tool build} on the target will check dependencies accurately,
       
   489   and rebuild outdated images on demand.
       
   490 
       
   491   \<^medskip> Option \<^verbatim>\<open>-H\<close> tells the underlying \<^verbatim>\<open>rsync\<close> process to purge the \<^verbatim>\<open>heaps\<close>
       
   492   directory on the target, before uploading new images from \<^verbatim>\<open>-I\<close> options. The
       
   493   default is to work monotonically: old material that is not overwritten
       
   494   remains unchanged, it is never deleted. Over time, this may lead to
       
   495   unreachable garbage, due to changes in session names or Poly/ML versions.
       
   496   Option \<^verbatim>\<open>-H\<close> helps to avoid wasting file-system space, although @{tool
       
   497   build} does not require it, due to its precise checking of all dependencies.
       
   498 
       
   499   \<^medskip> Option \<^verbatim>\<open>-S\<close> uses \<^verbatim>\<open>rsync --protect-args\<close>, but this requires at least
       
   500   \<^verbatim>\<open>rsync 3.0.0\<close>, while many versions of macOS still ship \<^verbatim>\<open>rsync 2.9.x\<close>.
       
   501   Since Isabelle + AFP don't use spaces or other special characters, it is
       
   502   usually safe to omit this non-portable option.
       
   503 \<close>
       
   504 
       
   505 subsubsection \<open>Examples\<close>
       
   506 
       
   507 text \<open>
       
   508   Quick testing of Isabelle + AFP on a remote machine: upload changed sources,
       
   509   jars, and local sessions images for \<^verbatim>\<open>HOL\<close>:
       
   510   @{verbatim [display] \<open>  isabelle sync -A: -I HOL -J testmachine:test/isabelle_afp\<close>}
       
   511   Assuming that the local \<^verbatim>\<open>HOL\<close> hierarchy has been up-to-date, and the local
       
   512   and remote ML platforms coincide, a remote @{tool build} will proceed
       
   513   without building \<^verbatim>\<open>HOL\<close> again.
       
   514 
       
   515   \<^medskip> A variation for accurate testing of Isabelle + AFP: no option \<^verbatim>\<open>-J\<close>, but
       
   516   option \<^verbatim>\<open>-T\<close> to disable the default ``quick check'' of \<^verbatim>\<open>rsync\<close> (which only
       
   517   inspects file sizes and date stamps):
       
   518   @{verbatim [display] \<open> isabelle sync -A: -T testmachine:test/isabelle_afp\<close>}
       
   519 \<close>
       
   520 
       
   521 
   428 section \<open>Output the version identifier of the Isabelle distribution\<close>
   522 section \<open>Output the version identifier of the Isabelle distribution\<close>
   429 
   523 
   430 text \<open>
   524 text \<open>
   431   The @{tool_def version} tool displays Isabelle version information:
   525   The @{tool_def version} tool displays Isabelle version information:
   432   @{verbatim [display]
   526   @{verbatim [display]