src/Doc/Isar_Ref/Document_Preparation.thy
changeset 69962 82e945d472d5
parent 69597 ff784d5a5bfb
child 70121 61e26527480e
equal deleted inserted replaced
69961:708743578e45 69962:82e945d472d5
   464     (\secref{sec:antiq}), usually with proper nesting of sub-languages via
   464     (\secref{sec:antiq}), usually with proper nesting of sub-languages via
   465     text cartouches.
   465     text cartouches.
   466 \<close>
   466 \<close>
   467 
   467 
   468 
   468 
   469 section \<open>Markup via command tags \label{sec:tags}\<close>
   469 section \<open>Document markers and command tags \label{sec:document-markers}\<close>
   470 
   470 
   471 text \<open>
   471 text \<open>
   472   Each Isabelle/Isar command may be decorated by additional presentation tags,
   472   \emph{Document markers} are formal comments of the form \<open>\<^marker>\<open>marker_body\<close>\<close>
   473   to indicate some modification in the way it is printed in the document.
   473   (using the control symbol \<^verbatim>\<open>\<^marker>\<close>) and may occur anywhere within the
       
   474   outer syntax of a command: the inner syntax of a marker body resembles that
       
   475   for attributes (\secref{sec:syn-att}). In contrast, \emph{Command tags} may
       
   476   only occur after a command keyword and are treated as special markers as
       
   477   explained below.
   474 
   478 
   475   \<^rail>\<open>
   479   \<^rail>\<open>
   476     @{syntax_def tags}: ( tag * )
   480     @{syntax_def marker}: '\<^marker>' @{syntax cartouche}
       
   481     ;
       
   482     @{syntax_def marker_body}: (@{syntax name} @{syntax args} * ',')
       
   483     ;
       
   484     @{syntax_def tags}: tag*
   477     ;
   485     ;
   478     tag: '%' (@{syntax short_ident} | @{syntax string})
   486     tag: '%' (@{syntax short_ident} | @{syntax string})
   479   \<close>
   487   \<close>
   480 
   488 
   481   Some tags are pre-declared for certain classes of commands, serving as
   489   Document markers are stripped from the document output, but surrounding
   482   default markup if no tags are given in the text:
   490   white-space is preserved: e.g.\ a marker at the end of a line does not
   483 
   491   affect the subsequent line break. Markers operate within the semantic
   484   \<^medskip>
   492   presentation context of a command, and may modify it to change the overall
   485   \begin{tabular}{ll}
   493   appearance of a command span (e.g.\ by adding tags).
   486     \<open>document\<close> & document markup commands \\
   494 
   487     \<open>theory\<close> & theory begin/end \\
   495   Each document marker has its own syntax defined in the theory context; the
   488     \<open>proof\<close> & all proof commands \\
   496   following markers are predefined in Isabelle/Pure:
   489     \<open>ML\<close> & all commands involving ML code \\
   497 
   490   \end{tabular}
   498   \<^rail>\<open>
   491   \<^medskip>
   499     (@@{document_marker_def title} |
       
   500      @@{document_marker_def creator} |
       
   501      @@{document_marker_def contributor} |
       
   502      @@{document_marker_def date} |
       
   503      @@{document_marker_def license} |
       
   504      @@{document_marker_def description}) @{syntax embedded}
       
   505     ;
       
   506     @@{document_marker_def tag} @{syntax name}
       
   507   \<close>
       
   508 
       
   509     \<^descr> \<open>\<^marker>\<open>title arg\<close>\<close>, \<open>\<^marker>\<open>creator arg\<close>\<close>, \<open>\<^marker>\<open>contributor arg\<close>\<close>, \<open>\<^marker>\<open>date arg\<close>\<close>,
       
   510     \<open>\<^marker>\<open>license arg\<close>\<close>, and \<open>\<^marker>\<open>description arg\<close>\<close> produce markup in the PIDE
       
   511     document, without any immediate effect on typesetting. This vocabulary is
       
   512     taken from the Dublin Core Metadata
       
   513     Initiative\<^footnote>\<open>\<^url>\<open>https://www.dublincore.org/specifications/dublin-core/dcmi-terms\<close>\<close>.
       
   514     The argument is an uninterpreted string, except for @{document_marker
       
   515     description}, which consists of words that are subject to spell-checking.
       
   516 
       
   517     \<^descr> \<open>\<^marker>\<open>tag name\<close>\<close> updates the list of command tags in the presentation
       
   518     context: later declarations take precedence, so \<open>\<^marker>\<open>tag a, tag b, tag c\<close>\<close>
       
   519     produces a reversed list. The default tags are given by the original
       
   520     \<^theory_text>\<open>keywords\<close> declaration of a command, and the system option
       
   521     @{system_option_ref document_tags}.
       
   522 
       
   523     An old-style command tag \<^verbatim>\<open>%\<close>\<open>name\<close> is treated like a document marker
       
   524     \<open>\<^marker>\<open>tag name\<close>\<close>: the list of command tags precedes the list of document
       
   525     markers. The head of the resulting tags in the presentation context is
       
   526     turned into {\LaTeX} environments to modify the type-setting. The
       
   527     following tags are pre-declared for certain classes of commands, and serve
       
   528     as default markup for certain kinds of commands:
       
   529 
       
   530     \<^medskip>
       
   531     \begin{tabular}{ll}
       
   532       \<open>document\<close> & document markup commands \\
       
   533       \<open>theory\<close> & theory begin/end \\
       
   534       \<open>proof\<close> & all proof commands \\
       
   535       \<open>ML\<close> & all commands involving ML code \\
       
   536     \end{tabular}
       
   537     \<^medskip>
   492 
   538 
   493   The Isabelle document preparation system @{cite "isabelle-system"} allows
   539   The Isabelle document preparation system @{cite "isabelle-system"} allows
   494   tagged command regions to be presented specifically, e.g.\ to fold proof
   540   tagged command regions to be presented specifically, e.g.\ to fold proof
   495   texts, or drop parts of the text completely.
   541   texts, or drop parts of the text completely.
   496 
   542 
   497   For example ``@{command "by"}~\<open>%invisible auto\<close>'' causes that piece of proof
   543   For example ``\<^theory_text>\<open>by auto\<close>~\<open>\<^marker>\<open>tag invisible\<close>\<close>'' causes that piece of proof to
   498   to be treated as \<open>invisible\<close> instead of \<open>proof\<close> (the default), which may be
   544   be treated as \<open>invisible\<close> instead of \<open>proof\<close> (the default), which may be
   499   shown or hidden depending on the document setup. In contrast, ``@{command
   545   shown or hidden depending on the document setup. In contrast, ``\<^theory_text>\<open>by
   500   "by"}~\<open>%visible auto\<close>'' forces this text to be shown invariably.
   546   auto\<close>~\<open>\<^marker>\<open>tag visible\<close>\<close>'' forces this text to be shown invariably.
   501 
   547 
   502   Explicit tag specifications within a proof apply to all subsequent commands
   548   Explicit tag specifications within a proof apply to all subsequent commands
   503   of the same level of nesting. For example, ``@{command "proof"}~\<open>%visible
   549   of the same level of nesting. For example, ``\<^theory_text>\<open>proof\<close>~\<open>\<^marker>\<open>tag invisible\<close>
   504   \<dots>\<close>~@{command "qed"}'' forces the whole sub-proof to be typeset as \<open>visible\<close>
   550   \<dots>\<close>~\<^theory_text>\<open>qed\<close>'' forces the whole sub-proof to be typeset as \<open>visible\<close> (unless
   505   (unless some of its parts are tagged differently).
   551   some of its parts are tagged differently).
   506 
   552 
   507   \<^medskip>
   553   \<^medskip>
   508   Command tags merely produce certain markup environments for type-setting.
   554   Command tags merely produce certain markup environments for type-setting.
   509   The meaning of these is determined by {\LaTeX} macros, as defined in
   555   The meaning of these is determined by {\LaTeX} macros, as defined in
   510   \<^file>\<open>~~/lib/texinputs/isabelle.sty\<close> or by the document author. The Isabelle
   556   \<^file>\<open>~~/lib/texinputs/isabelle.sty\<close> or by the document author. The Isabelle