# HG changeset patch # User wenzelm # Date 1546715522 -3600 # Node ID caa7e406056d7627d5ec2d04a0b66ad081420321 # Parent 81caae4fc4fa43fae3a28853bd95395c4d6cab41 documentation on "isabelle update"; diff -r 81caae4fc4fa -r caa7e406056d src/Doc/Isar_Ref/document/style.sty --- a/src/Doc/Isar_Ref/document/style.sty Sat Jan 05 17:33:20 2019 +0100 +++ b/src/Doc/Isar_Ref/document/style.sty Sat Jan 05 20:12:02 2019 +0100 @@ -17,6 +17,7 @@ %% Isar \newcommand{\isasymBBAR}{{\,\newdimen{\tmpheight}\settoheight\tmpheight{\isacharbar}\rule{1pt}{\tmpheight}\,}} \renewcommand{\isacommand}[1]{\isakeyword{\isadigitreset#1}} +\newcommand{\isasymdoublequote}{\texttt{\upshape"}} %% ML \newenvironment{mldecls}{\par\noindent\begingroup\def\isanewline{\\}\begin{tabular}{ll}}{\end{tabular}\medskip\endgroup} diff -r 81caae4fc4fa -r caa7e406056d src/Doc/System/Sessions.thy --- a/src/Doc/System/Sessions.thy Sat Jan 05 17:33:20 2019 +0100 +++ b/src/Doc/System/Sessions.thy Sat Jan 05 20:12:02 2019 +0100 @@ -674,7 +674,106 @@ \ISABELLE_TOOL_JAVA_OPTIONS="-Xms4g -Xmx32g -Xss16m" ML_OPTIONS="--minheap 4G --maxheap 32G" \} +\ + +section \Update theory sources based on PIDE markup \label{sec:tool-update}\ + +text \ + The @{tool_def "update"} tool updates theory sources based on markup that is + produced from a running PIDE session (similar to @{tool dump} + \secref{sec:tool-dump}). Its command-line usage is: @{verbatim [display] +\Usage: isabelle update [OPTIONS] [SESSIONS ...] + + Options are: + -B NAME include session NAME and all descendants + -D DIR include session directory and select its sessions + -R operate on requirements of selected sessions + -X NAME exclude sessions from group NAME and all descendants + -a select all sessions + -d DIR include session directory + -g NAME select session group NAME + -l NAME logic session name (default ISABELLE_LOGIC="HOL") + -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) + -s system build mode for logic image + -u OPT overide update option: shortcut for "-o update_OPT" + -v verbose + -x NAME exclude session NAME and all descendants + + Update theory sources based on PIDE markup.\} + + \<^medskip> Options \<^verbatim>\-B\, \<^verbatim>\-D\, \<^verbatim>\-R\, \<^verbatim>\-X\, \<^verbatim>\-a\, \<^verbatim>\-d\, \<^verbatim>\-g\, \<^verbatim>\-x\ and the + remaining command-line arguments specify sessions as in @{tool build} + (\secref{sec:tool-build}) or @{tool dump} (\secref{sec:tool-dump}). + + \<^medskip> Options \<^verbatim>\-l\ and \<^verbatim>\-s\ specify the underlying logic image is in @{tool + dump} (\secref{sec:tool-dump}). + + \<^medskip> Option \<^verbatim>\-v\ increases the general level of verbosity. + + \<^medskip> Option \<^verbatim>\-o\ overrides Isabelle system options as for @{tool build} + (\secref{sec:tool-build}). Option \<^verbatim>\-u\ refers to specific \<^verbatim>\update\ + options, by relying on naming convention: ``\<^verbatim>\-u\~\OPT\'' is a shortcut for + ``\<^verbatim>\-o\~\<^verbatim>\update_\\OPT\''. + + \<^medskip> The following update options are supported: + + \<^item> @{system_option update_inner_syntax_cartouches} to update inner syntax + (types, terms, etc.)~to use cartouches, instead of double-quoted strings + or atomic identifiers. For example, ``\<^theory_text>\lemma \x = + x\\'' is replaced by ``\<^theory_text>\lemma \x = x\\'', and ``\<^theory_text>\fix x\'' + is replaced by ``\<^theory_text>\fix \x\\''. + + \<^item> @{system_option update_mixfix_cartouches} to update mixfix templates to + use cartouches instead of double-quoted strings. For example, ``\<^theory_text>\(infixl + \+\ 65)\'' is replaced by ``\<^theory_text>\(infixl \+\ + 65)\''. + + \<^item> @{system_option update_control_cartouches} to update antiquotations to + use the compact form with control symbol and cartouche argument. For + example, ``\@{term \x + y\}\'' is replaced by + ``\\<^term>\x + y\\'' (the control symbol is literally \<^verbatim>\\<^term>\.) + + It is also possible to produce custom updates in Isabelle/ML, by reporting + \<^ML>\Markup.update\ with the precise source position and a replacement + text. This operation should be made conditional on specific system options, + similar to the ones above. Searching the above option names in ML sources of + \<^dir>\$ISABELLE_HOME/src/Pure\ provides some examples. + + Different update options can be in conflict by producing overlapping edits: + this may require to run @{tool update} multiple times, but it is often + better to enable particular update options separately and commit the changes + one-by-one. +\ + + +subsubsection \Examples\ + +text \ + Update some cartouche notation in all theory sources required for session + \<^verbatim>\HOL-Analysis\: + + @{verbatim [display] \isabelle update -u mixfix_cartouches -l Pure HOL-Analysis\} + + \<^smallskip> Update the same for all application sessions based on \<^verbatim>\HOL-Analysis\ --- + its image is taken as starting point and its sources are not touched: + + @{verbatim [display] \isabelle update -u mixfix_cartouches -l HOL-Analysis -B HOL-Analysis\} + + \<^smallskip> This two-stage approach reduces resource requirements of the running PIDE + session: a base image like \<^verbatim>\HOL-Analysis\ (or \<^verbatim>\HOL\, \<^verbatim>\HOL-Library\) is + more compact than importing all required theory (and ML) source files from + \<^verbatim>\Pure\. + + \<^smallskip> Update sessions that build on \<^verbatim>\HOL-Proofs\, which need to be run + separately with special options as follows: + + @{verbatim [display] \isabelle update -u mixfix_cartouches -l HOL-Proofs -B HOL-Proofs + -o record_proofs=2 -o parallel_proofs=0\} + + \<^smallskip> See also the end of \secref{sec:tool-dump} for hints on increasing + Isabelle/ML heap sizes for very big PIDE processes that include many + sessions, notably from the Archive of Formal Proofs. \ end