\chapter{Syntax primitives}The rather generic framework of Isabelle/Isar syntax emerges from three mainsyntactic categories: \emph{commands} of the top-level Isar engine (coveringtheory and proof elements), \emph{methods} for general goal refinements(analogous to traditional ``tactics''), and \emph{attributes} for operationson facts (within a certain context). Here we give a reference of basicsyntactic entities underlying Isabelle/Isar syntax in a bottom-up manner.Concrete theory and proof language elements will be introduced later on.\medskipIn order to get started with writing well-formed Isabelle/Isar documents, themost important aspect to be noted is the difference of \emph{inner} versus\emph{outer} syntax. Inner syntax is that of Isabelle types and terms of thelogic, while outer syntax is that of Isabelle/Isar theory sources (includingproofs). As a general rule, inner syntax entities may occur only as\emph{atomic entities} within outer syntax. For example, the string\texttt{"x + y"} and identifier \texttt{z} are legal term specificationswithin a theory, while \texttt{x + y} is not.\begin{warn} Old-style Isabelle theories used to fake parts of the inner syntax of types, with rather complicated rules when quotes may be omitted. Despite the minor drawback of requiring quotes more often, the syntax of Isabelle/Isar is somewhat simpler and more robust in that respect.\end{warn}Printed theory documents usually omit quotes to gain readability (this is amatter of {\LaTeX} macro setup, say via \verb,\isabellestyle,, see also\cite{isabelle-sys}). Experienced users of Isabelle/Isar may easilyreconstruct the lost technical information, while mere readers need not careabout quotes at all.\medskipIsabelle/Isar input may contain any number of input termination characters``\texttt{;}'' (semicolon) to separate commands explicitly. This isparticularly useful in interactive shell sessions to make clear where thecurrent command is intended to end. Otherwise, the interpreter loop willcontinue to issue a secondary prompt ``\verb,#,'' until an end-of-command isclearly recognized from the input syntax, e.g.\ encounter of the next commandkeyword.Advanced interfaces such as Proof~General \cite{proofgeneral} do not requireexplicit semicolons, the amount of input text is determined automatically byinspecting the present content of the Emacs text buffer. In the printedpresentation of Isabelle/Isar documents semicolons are omitted altogether forreadability.\begin{warn} Proof~General requires certain syntax classification tables in order to achieve properly synchronized interaction with the Isabelle/Isar process. These tables need to be consistent with the Isabelle version and particular logic image to be used in a running session (common object-logics may well change the outer syntax). The standard setup should work correctly with any of the ``official'' logic images derived from Isabelle/HOL (including HOLCF etc.). Users of alternative logics may need to tell Proof~General explicitly, e.g.\ by giving an option \verb,-k ZF, (in conjunction with \verb,-l ZF, to specify the default logic image).\end{warn}\section{Lexical matters}\label{sec:lex-syntax}The Isabelle/Isar outer syntax provides token classes as presented below; mostof these coincide with the inner lexical syntax as presented in\cite{isabelle-ref}.\indexoutertoken{ident}\indexoutertoken{longident}\indexoutertoken{symident}\indexoutertoken{nat}\indexoutertoken{var}\indexoutertoken{typefree}\indexoutertoken{typevar}\indexoutertoken{string}\indexoutertoken{altstring}\indexoutertoken{verbatim}\begin{matharray}{rcl} ident & = & letter\,quasiletter^* \\ longident & = & ident (\verb,.,ident)^+ \\ symident & = & sym^+ ~|~ \verb,\<,ident\verb,>, \\ nat & = & digit^+ \\ var & = & ident ~|~ \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\ typefree & = & \verb,',ident \\ typevar & = & typefree ~|~ \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\ string & = & \verb,", ~\dots~ \verb,", \\ altstring & = & \backquote ~\dots~ \backquote \\ verbatim & = & \verb,{*, ~\dots~ \verb,*}, \\[1ex] letter & = & latin ~|~ \verb,\<,latin\verb,>, ~|~ \verb,\<,latin\,latin\verb,>, ~|~ greek ~|~ \\ & & \verb,\<^isub>, ~|~ \verb,\<^isup>, \\ quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\ latin & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\ digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\ sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~ %$ \verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \verb,:, ~|~ \\ & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~ \verb,^, ~|~ \verb,_, ~|~ \verb,|, ~|~ \verb,~, \\greek & = & \verb,\<alpha>, ~|~ \verb,\<beta>, ~|~ \verb,\<gamma>, ~|~ \verb,\<delta>, ~| \\ & & \verb,\<epsilon>, ~|~ \verb,\<zeta>, ~|~ \verb,\<eta>, ~|~ \verb,\<theta>, ~| \\ & & \verb,\<iota>, ~|~ \verb,\<kappa>, ~|~ \verb,\<mu>, ~|~ \verb,\<nu>, ~| \\ & & \verb,\<xi>, ~|~ \verb,\<pi>, ~|~ \verb,\<rho>, ~|~ \verb,\<sigma>, ~| \\ & & \verb,\<tau>, ~|~ \verb,\<upsilon>, ~|~ \verb,\<phi>, ~|~ \verb,\<psi>, ~| \\ & & \verb,\<omega>, ~|~ \verb,\<Gamma>, ~|~ \verb,\<Delta>, ~|~ \verb,\<Theta>, ~| \\ & & \verb,\<Lambda>, ~|~ \verb,\<Xi>, ~|~ \verb,\<Pi>, ~|~ \verb,\<Sigma>, ~| \\ & & \verb,\<Upsilon>, ~|~ \verb,\<Phi>, ~|~ \verb,\<Psi>, ~|~ \verb,\<Omega>, \\\end{matharray}The syntax of $string$ admits any characters, including newlines; ``\verb|"|''(double-quote) and ``\verb|\|'' (backslash) need to be escaped by a backslash.Alternative strings according to $altstring$ are analogous, using singleback-quotes instead. The body of $verbatim$ may consist of any text notcontaining ``\verb|*}|''; this allows convenient inclusion of quotes withoutfurther escapes. The greek letters do \emph{not} include \verb,\<lambda>,,which is already used differently in the meta-logic.Common mathematical symbols such as $\forall$ are represented in Isabelle as\verb,\<forall>,. There are infinitely many legal symbols like this, althoughproper presentation is left to front-end tools such as {\LaTeX} orProof~General with the X-Symbol package. A list of standard Isabelle symbolsthat work well with these tools is given in \cite[appendix~A]{isabelle-sys}.Comments take the form \texttt{(*~\dots~*)} and may be nested, althoughuser-interface tools may prevent this. Note that \texttt{(*~\dots~*)}indicate source comments only, which are stripped after lexical analysis ofthe input. The Isar document syntax also provides formal comments that areconsidered as part of the text (see \S\ref{sec:comments}).\begin{warn} Proof~General does not handle nested comments properly; it is also unable to keep \verb,(*,\,/\,\verb,{*, and \verb,*),\,/\,\verb,*}, apart, despite their rather different meaning. These are inherent problems of Emacs legacy. Users should not be overly aggressive about nesting or alternating these delimiters.\end{warn}\section{Common syntax entities}Subsequently, we introduce several basic syntactic entities, such as names,terms, and theorem specifications, which have been factored out of the actualIsar language elements to be described later.Note that some of the basic syntactic entities introduced below (e.g.\\railqtok{name}) act much like tokens rather than plain nonterminals (e.g.\\railnonterm{sort}), especially for the sake of error messages. E.g.\ syntaxelements like $\CONSTS$ referring to \railqtok{name} or \railqtok{type} wouldreally report a missing name or type rather than any of the constituentprimitive tokens such as \railtok{ident} or \railtok{string}.\subsection{Names}Entity \railqtok{name} usually refers to any name of types, constants,theorems etc.\ that are to be \emph{declared} or \emph{defined} (so qualifiedidentifiers are excluded here). Quoted strings provide an escape fornon-identifier names or those ruled out by outer syntax keywords (e.g.\\verb|"let"|). Already existing objects are usually referenced by\railqtok{nameref}.\indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}\indexoutertoken{int}\begin{rail} name: ident | symident | string | nat ; parname: '(' name ')' ; nameref: name | longident ; int: nat | '-' nat ;\end{rail}\subsection{Comments}\label{sec:comments}Large chunks of plain \railqtok{text} are usually given \railtok{verbatim},i.e.\ enclosed in \verb|{*|~\dots~\verb|*}|. For convenience, any of thesmaller text units conforming to \railqtok{nameref} are admitted as well. Amarginal \railnonterm{comment} is of the form \texttt{--} \railqtok{text}.Any number of these may occur within Isabelle/Isar commands.\indexoutertoken{text}\indexouternonterm{comment}\begin{rail} text: verbatim | nameref ; comment: '--' text ;\end{rail}\subsection{Type classes, sorts and arities}Classes are specified by plain names. Sorts have a very simple inner syntax,which is either a single class name $c$ or a list $\{c@1, \dots, c@n\}$referring to the intersection of these classes. The syntax of type arities isgiven directly at the outer level.\railalias{subseteq}{\isasymsubseteq}\railterm{subseteq}\indexouternonterm{sort}\indexouternonterm{arity}\indexouternonterm{classdecl}\begin{rail} classdecl: name (('<' | subseteq) (nameref + ','))? ; sort: nameref ; arity: ('(' (sort + ',') ')')? sort ;\end{rail}\subsection{Types and terms}\label{sec:types-terms}The actual inner Isabelle syntax, that of types and terms of the logic, is fartoo sophisticated in order to be modelled explicitly at the outer theorylevel. Basically, any such entity has to be quoted to turn it into a singletoken (the parsing and type-checking is performed internally later). Forconvenience, a slightly more liberal convention is adopted: quotes may beomitted for any type or term that is already atomic at the outer level. Forexample, one may just write \texttt{x} instead of \texttt{"x"}. Note thatsymbolic identifiers (e.g.\ \texttt{++} or $\forall$) are available as well,provided these have not been superseded by commands or other keywords already(e.g.\ \texttt{=} or \texttt{+}).\indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}\begin{rail} type: nameref | typefree | typevar ; term: nameref | var ; prop: term ;\end{rail}Positional instantiations are indicated by giving a sequence of terms, or theplaceholder ``$\_$'' (underscore), which means to skip a position.\indexoutertoken{inst}\indexoutertoken{insts}\begin{rail} inst: underscore | term ; insts: (inst *) ;\end{rail}Type declarations and definitions usually refer to \railnonterm{typespec} onthe left-hand side. This models basic type constructor application at theouter syntax level. Note that only plain postfix notation is available here,but no infixes.\indexouternonterm{typespec}\begin{rail} typespec: (() | typefree | '(' ( typefree + ',' ) ')') name ;\end{rail}\subsection{Mixfix annotations}Mixfix annotations specify concrete \emph{inner} syntax of Isabelle types andterms. Some commands such as $\TYPES$ (see \S\ref{sec:types-pure}) admitinfixes only, while $\CONSTS$ (see \S\ref{sec:consts}) and$\isarkeyword{syntax}$ (see \S\ref{sec:syn-trans}) support the full range ofgeneral mixfixes and binders.\indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}\begin{rail} infix: '(' ('infix' | 'infixl' | 'infixr') string? nat ')' ; mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')' ; structmixfix: mixfix | '(' 'structure' ')' ; prios: '[' (nat + ',') ']' ;\end{rail}Here the \railtok{string} specifications refer to the actual mixfix template(see also \cite{isabelle-ref}), which may include literal text, spacing,blocks, and arguments (denoted by ``$_$''); the special symbol \verb,\<index>,(printed as ``\i'') represents an index argument that specifies an implicitstructure reference (see also \S\ref{sec:locale}). Infix and binderdeclarations provide common abbreviations for particular mixfix declarations.So in practice, mixfix templates mostly degenerate to literal text forconcrete syntax, such as ``\verb,++,'' for an infix symbol, or ``\verb,++,\i''for an infix of an implicit structure.\subsection{Proof methods}\label{sec:syn-meth}Proof methods are either basic ones, or expressions composed ofmethods via ``\texttt{,}'' (sequential composition), ``\texttt{|}''(alternative choices), ``\texttt{?}'' (try), ``\texttt{+}'' (repeat atleast once), ``\texttt{[$n$]}'' (restriction to first $n$ sub-goals,default $n = 1$). In practice, proof methods are usually just a commaseparated list of \railqtok{nameref}~\railnonterm{args}specifications. Note that parentheses may be dropped for singlemethod specifications (with no arguments).\indexouternonterm{method}\begin{rail} method: (nameref | '(' methods ')') (() | '?' | '+' | '[' nat? ']') ; methods: (nameref args | method) + (',' | '|') ;\end{rail}Proper Isar proof methods do \emph{not} admit arbitrary goaladdressing, but refer either to the first sub-goal or all sub-goalsuniformly. The goal restriction operator ``\texttt{[$n$]}'' evaluatesa method expression within a sandbox consisting of the first $n$sub-goals (which need to exist). For example,$simp_all\mbox{\tt[}3\mbox{\tt]}$ simplifies the first threesub-goals, while $(rule~foo, simp_all)\mbox{\tt[]}$ simplifies all newgoals that emerge from applying rule $foo$ to the originally firstone.Improper methods, notably tactic emulations, offer a separatelow-level goal addressing scheme as explicit argument to theindividual tactic being involved. Here $[!]$ refers to all goals, and$[n-]$ to all goals starting from $n$,\indexouternonterm{goalspec}\begin{rail} goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']' ;\end{rail}\subsection{Attributes and theorems}\label{sec:syn-att}Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their own``semi-inner'' syntax, in the sense that input conforming to\railnonterm{args} below is parsed by the attribute a second time. Theattribute argument specifications may be any sequence of atomic entities(identifiers, strings etc.), or properly bracketed argument lists. Below\railqtok{atom} refers to any atomic entity, including any \railtok{keyword}conforming to \railtok{symident}.\indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}\begin{rail} atom: nameref | typefree | typevar | var | nat | keyword ; arg: atom | '(' args ')' | '[' args ']' ; args: arg * ; attributes: '[' (nameref args * ',') ']' ;\end{rail}Theorem specifications come in several flavors: \railnonterm{axmdecl} and\railnonterm{thmdecl} usually refer to axioms, assumptions or results of goalstatements, while \railnonterm{thmdef} collects lists of existing theorems.Existing theorems are given by \railnonterm{thmref} and \railnonterm{thmrefs},the former requires an actual singleton result. There are three forms oftheorem references: (1) named facts $a$, (2) selections from named facts $a(i,j - k)$, or (3) literal fact propositions using $altstring$ syntax$\backquote\phi\backquote$, (see also method $fact$ in\S\ref{sec:pure-meth-att}).Any kind of theorem specification may include lists of attributes both on theleft and right hand sides; attributes are applied to any immediately precedingfact. If names are omitted, the theorems are not stored within the theoremdatabase of the theory or proof context, but any given attributes are appliednonetheless.\indexouternonterm{axmdecl}\indexouternonterm{thmdecl}\indexouternonterm{thmdef}\indexouternonterm{thmref}\indexouternonterm{thmrefs}\indexouternonterm{selection}\begin{rail} axmdecl: name attributes? ':' ; thmdecl: thmbind ':' ; thmdef: thmbind '=' ; thmref: (nameref selection? | altstring) attributes? ; thmrefs: thmref + ; thmbind: name attributes | name | attributes ; selection: '(' ((nat | nat '-' nat?) + ',') ')' ;\end{rail}\subsection{Term patterns and declarations}\label{sec:term-decls}Wherever explicit propositions (or term fragments) occur in a proof text,casual binding of schematic term variables may be given specified via patternsof the form ``$\ISS{p@1\;\dots}{p@n}$''. There are separate versionsavailable for \railqtok{term}s and \railqtok{prop}s. The latter provides a$\CONCLNAME$ part with patterns referring the (atomic) conclusion of a rule.\indexouternonterm{termpat}\indexouternonterm{proppat}\begin{rail} termpat: '(' ('is' term +) ')' ; proppat: '(' ('is' prop +) ')' ;\end{rail}Declarations of local variables $x :: \tau$ and logical propositions $a :\phi$ represent different views on the same principle of introducing a localscope. In practice, one may usually omit the typing of $vars$ (due totype-inference), and the naming of propositions (due to implicit references ofcurrent facts). In any case, Isar proof elements usually admit to introducemultiple such items simultaneously.\indexouternonterm{vars}\indexouternonterm{props}\begin{rail} vars: (name+) ('::' type)? ; props: thmdecl? (prop proppat? +) ;\end{rail}The treatment of multiple declarations corresponds to the complementary focusof $vars$ versus $props$: in ``$x@1~\dots~x@n :: \tau$'' the typing refers toall variables, while in $a\colon \phi@1~\dots~\phi@n$ the naming refers to allpropositions collectively. Isar language elements that refer to $vars$ or$props$ typically admit separate typings or namings via another level ofiteration, with explicit $\AND$ separators; e.g.\ see $\FIXNAME$ and$\ASSUMENAME$ in \S\ref{sec:proof-context}.\subsection{Antiquotations}\label{sec:antiq}\begin{matharray}{rcl} thm & : & \isarantiq \\ prop & : & \isarantiq \\ term & : & \isarantiq \\ const & : & \isarantiq \\ typeof & : & \isarantiq \\ typ & : & \isarantiq \\ thm_style & : & \isarantiq \\ term_style & : & \isarantiq \\ text & : & \isarantiq \\ goals & : & \isarantiq \\ subgoals & : & \isarantiq \\ prf & : & \isarantiq \\ full_prf & : & \isarantiq \\ ML & : & \isarantiq \\ ML_type & : & \isarantiq \\ ML_struct & : & \isarantiq \\\end{matharray}The text body of formal comments (see also \S\ref{sec:comments}) may containantiquotations of logical entities, such as theorems, terms and types, whichare to be presented in the final output produced by the Isabelle documentpreparation system (see also \S\ref{sec:document-prep}).Thus embedding of``\texttt{{\at}{\ttlbrace}term~[show_types]~"f(x)~=~a~+~x"{\ttrbrace}}''within a text block would cause\isa{(f{\isasymColon}'a~{\isasymRightarrow}~'a)~(x{\isasymColon}'a)~=~(a{\isasymColon}'a)~+~x}to appear in the final {\LaTeX} document. Also note that theoremantiquotations may involve attributes as well. For example,\texttt{{\at}{\ttlbrace}thm~sym~[no_vars]{\ttrbrace}} would print thestatement where all schematic variables have been replaced by fixed ones,which are easier to read.\indexisarant{thm}\indexisarant{prop}\indexisarant{term}\indexisarant{const}\indexisarant{typeof}\indexisarant{typ}\indexisarant{thm-style}\indexisarant{term-style}\indexisarant{text}\indexisarant{goals}\indexisarant{subgoals}\indexisarant{prf}\indexisarant{full-prf}\indexisarant{ML}\indexisarant{ML-type}\indexisarant{ML-struct}\begin{rail} atsign lbrace antiquotation rbrace ; antiquotation: 'thm' options thmrefs | 'prop' options prop | 'term' options term | 'const' options term | 'typeof' options term | 'typ' options type | 'thm\_style' options name thmref | 'term\_style' options name term | 'text' options name | 'goals' options | 'subgoals' options | 'prf' options thmrefs | 'full\_prf' options thmrefs | 'ML' options name | 'ML\_type' options name | 'ML\_struct' options name ; options: '[' (option * ',') ']' ; option: name | name '=' name ;\end{rail}Note that the syntax of antiquotations may \emph{not} include source comments\texttt{(*~\dots~*)} or verbatim text \verb|{*|~\dots~\verb|*}|.\begin{descr}\item [$\at\{thm~\vec a\}$] prints theorems $\vec a$. Note that attribute specifications may be included as well (see also \S\ref{sec:syn-att}); the $no_vars$ operation (see \S\ref{sec:misc-meth-att}) would be particularly useful to suppress printing of schematic variables.\item [$\at\{prop~\phi\}$] prints a well-typed proposition $\phi$.\item [$\at\{term~t\}$] prints a well-typed term $t$.\item [$\at\{const~c\}$] prints a well-defined constant $c$.\item [$\at\{typeof~t\}$] prints the type of a well-typed term $t$.\item [$\at\{typ~\tau\}$] prints a well-formed type $\tau$.\item [$\at\{thm_style~s~a\}$] prints theorem $a$, previously applying a style $s$ to it (see below).\item [$\at\{term_style~s~t\}$] prints a well-typed term $t$ after applying a style $s$ to it (see below).\item [$\at\{text~s\}$] prints uninterpreted source text $s$. This is particularly useful to print portions of text according to the Isabelle {\LaTeX} output style, without demanding well-formedness (e.g.\ small pieces of terms that should not be parsed or type-checked yet).\item [$\at\{goals\}$] prints the current \emph{dynamic} goal state. This is mainly for support of tactic-emulation scripts within Isar --- presentation of goal states does not conform to actual human-readable proof documents. Please do not include goal states into document output unless you really know what you are doing!\item [$\at\{subgoals\}$] is similar to $goals$, but does not print the main goal.\item [$\at\{prf~\vec a\}$] prints the (compact) proof terms corresponding to the theorems $\vec a$. Note that this requires proof terms to be switched on for the current object logic (see the ``Proof terms'' section of the Isabelle reference manual for information on how to do this).\item [$\at\{full_prf~\vec a\}$] is like $\at\{prf~\vec a\}$, but displays the full proof terms, i.e.\ also displays information omitted in the compact proof term, which is denoted by ``$_$'' placeholders there.\item [$\at\{ML~s\}$, $\at\{ML_type~s\}$, and $\at\{ML_struct~s\}$] check text $s$ as ML value, type, and structure, respectively. If successful, the source is displayed verbatim.\end{descr}\medskipThe following standard styles for use with $thm_style$ and $term_style$ areavailable:\begin{descr}\item [$lhs$] extracts the first argument of any application form with at least two arguments -- typically meta-level or object-level equality, or any other binary relation.\item [$rhs$] is like $lhs$, but extracts the second argument.\item [$concl$] extracts the conclusion $C$ from a nested meta-level implication $A@1 \Imp \cdots A@n \Imp C$.\item [$prem1$, \dots, $prem9$] extract premise number $1$, \dots, $9$, respectively, from a nested meta-level implication $A@1 \Imp \cdots A@n \Imp C$.\end{descr}\medskipThe following options are available to tune the output. Note that most ofthese coincide with ML flags of the same names (see also \cite{isabelle-ref}).\begin{descr}\item[$show_types = bool$ and $show_sorts = bool$] control printing of explicit type and sort constraints.\item[$show_structs = bool$] controls printing of implicit structures.\item[$long_names = bool$] forces names of types and constants etc.\ to be printed in their fully qualified internal form.\item[$short_names = bool$] forces names of types and constants etc.\ to be printed unqualified. Note that internalizing the output again in the current context may well yield a different result.\item[$unique_names = bool$] determines whether the printed version of qualified names should be made sufficiently long to avoid overlap with names declared further back. Set to $false$ for more concise output.\item[$eta_contract = bool$] prints terms in $\eta$-contracted form.\item[$display = bool$] indicates if the text is to be output as multi-line ``display material'', rather than a small piece of text without line breaks (which is the default).\item[$breaks = bool$] controls line breaks in non-display material.\item[$quotes = bool$] indicates if the output should be enclosed in double quotes.\item[$mode = name$] adds $name$ to the print mode to be used for presentation (see also \cite{isabelle-ref}). Note that the standard setup for {\LaTeX} output is already present by default, including the modes ``$latex$'', ``$xsymbols$'', ``$symbols$''.\item[$margin = nat$ and $indent = nat$] change the margin or indentation for pretty printing of display material.\item[$source = bool$] prints the source text of the antiquotation arguments, rather than the actual value. Note that this does not affect well-formedness checks of $thm$, $term$, etc. (only the $text$ antiquotation admits arbitrary output).\item[$goals_limit = nat$] determines the maximum number of goals to be printed.\item[$locale = name$] specifies an alternative context used for evaluating and printing the subsequent argument.\end{descr}For boolean flags, ``$name = true$'' may be abbreviated as ``$name$''. All ofthe above flags are disabled by default, unless changed from ML.\medskip Note that antiquotations do not only spare the author from tedioustyping of logical entities, but also achieve some degree ofconsistency-checking of informal explanations with formal developments:well-formedness of terms and types with respect to the current theory or proofcontext is ensured here.\subsection{Tagged commands}\label{sec:tags}Each Isabelle/Isar command may be decorated by presentation tags:\indexouternonterm{tags}\begin{rail} tags: ( tag * ) ; tag: '\%' (ident | string)\end{rail}The tags $theory$, $proof$, $ML$ are already pre-declared for certain classesof commands:\medskip\begin{tabular}{ll} $theory$ & theory begin and end \\ $proof$ & all proof commands \\ $ML$ & all commands involving ML code \\\end{tabular}\medskip The Isabelle document preparation system (see also\cite{isabelle-sys}) allows tagged command regions to be presentedspecifically, e.g.\ to fold proof texts, or drop parts of the text completely.For example ``$\BYNAME~\%invisible~(auto)$'' would cause that piece of proofto be treated as $invisible$ instead of $proof$ (the default), which may beeither show or hidden depending on the document setup. In contrast,``$\BYNAME~\%visible~(auto)$'' would force this text to be shown invariably.Explicit tag specifications within a proof apply to all subsequent commands ofthe same level of nesting. For example,``$\PROOFNAME~\%visible~\dots\QEDNAME$'' would force the whole sub-proof to betypeset as $visible$ (unless some of its parts are tagged differently).%%% Local Variables:%%% mode: latex%%% TeX-master: "isar-ref"%%% End: