doc-src/IsarRef/Thy/syntax.thy
changeset 26754 c0424e4de33d
child 26760 2de4ba348f06
equal deleted inserted replaced
26753:094d70c81243 26754:c0424e4de33d
       
     1 
       
     2 theory "syntax"
       
     3 imports CPure
       
     4 begin
       
     5 
       
     6 chapter {* Syntax primitives *}
       
     7 
       
     8 text {*
       
     9   The rather generic framework of Isabelle/Isar syntax emerges from
       
    10   three main syntactic categories: \emph{commands} of the top-level
       
    11   Isar engine (covering theory and proof elements), \emph{methods} for
       
    12   general goal refinements (analogous to traditional ``tactics''), and
       
    13   \emph{attributes} for operations on facts (within a certain
       
    14   context).  Subsequently we give a reference of basic syntactic
       
    15   entities underlying Isabelle/Isar syntax in a bottom-up manner.
       
    16   Concrete theory and proof language elements will be introduced later
       
    17   on.
       
    18 
       
    19   \medskip In order to get started with writing well-formed
       
    20   Isabelle/Isar documents, the most important aspect to be noted is
       
    21   the difference of \emph{inner} versus \emph{outer} syntax.  Inner
       
    22   syntax is that of Isabelle types and terms of the logic, while outer
       
    23   syntax is that of Isabelle/Isar theory sources (specifications and
       
    24   proofs).  As a general rule, inner syntax entities may occur only as
       
    25   \emph{atomic entities} within outer syntax.  For example, the string
       
    26   \texttt{"x + y"} and identifier \texttt{z} are legal term
       
    27   specifications within a theory, while \texttt{x + y} is not.
       
    28 
       
    29   Printed theory documents usually omit quotes to gain readability
       
    30   (this is a matter of {\LaTeX} macro setup, say via
       
    31   \verb,\isabellestyle,, see also \cite{isabelle-sys}).  Experienced
       
    32   users of Isabelle/Isar may easily reconstruct the lost technical
       
    33   information, while mere readers need not care about quotes at all.
       
    34 
       
    35   \medskip Isabelle/Isar input may contain any number of input
       
    36   termination characters ``\texttt{;}'' (semicolon) to separate
       
    37   commands explicitly.  This is particularly useful in interactive
       
    38   shell sessions to make clear where the current command is intended
       
    39   to end.  Otherwise, the interpreter loop will continue to issue a
       
    40   secondary prompt ``\verb,#,'' until an end-of-command is clearly
       
    41   recognized from the input syntax, e.g.\ encounter of the next
       
    42   command keyword.
       
    43 
       
    44   More advanced interfaces such as Proof~General \cite{proofgeneral}
       
    45   do not require explicit semicolons, the amount of input text is
       
    46   determined automatically by inspecting the present content of the
       
    47   Emacs text buffer.  In the printed presentation of Isabelle/Isar
       
    48   documents semicolons are omitted altogether for readability.
       
    49 
       
    50   \begin{warn}
       
    51     Proof~General requires certain syntax classification tables in
       
    52     order to achieve properly synchronized interaction with the
       
    53     Isabelle/Isar process.  These tables need to be consistent with
       
    54     the Isabelle version and particular logic image to be used in a
       
    55     running session (common object-logics may well change the outer
       
    56     syntax).  The standard setup should work correctly with any of the
       
    57     ``official'' logic images derived from Isabelle/HOL (including
       
    58     HOLCF etc.).  Users of alternative logics may need to tell
       
    59     Proof~General explicitly, e.g.\ by giving an option \verb,-k ZF,
       
    60     (in conjunction with \verb,-l ZF, to specify the default logic
       
    61     image).
       
    62   \end{warn}
       
    63 *}
       
    64 
       
    65 
       
    66 section {* Lexical matters \label{sec:lex-syntax} *}
       
    67 
       
    68 text {*
       
    69   The Isabelle/Isar outer syntax provides token classes as presented
       
    70   below; most of these coincide with the inner lexical syntax as
       
    71   presented in \cite{isabelle-ref}.
       
    72 
       
    73   \begin{matharray}{rcl}
       
    74     @{syntax_def ident} & = & letter\,quasiletter^* \\
       
    75     @{syntax_def longident} & = & ident (\verb,.,ident)^+ \\
       
    76     @{syntax_def symident} & = & sym^+ ~|~ \verb,\,\verb,<,ident\verb,>, \\
       
    77     @{syntax_def nat} & = & digit^+ \\
       
    78     @{syntax_def var} & = & ident ~|~ \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\
       
    79     @{syntax_def typefree} & = & \verb,',ident \\
       
    80     @{syntax_def typevar} & = & typefree ~|~ \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\
       
    81     @{syntax_def string} & = & \verb,", ~\dots~ \verb,", \\
       
    82     @{syntax_def altstring} & = & \backquote ~\dots~ \backquote \\
       
    83     @{syntax_def verbatim} & = & \verb,{*, ~\dots~ \verb,*,\verb,}, \\[1ex]
       
    84 
       
    85     letter & = & latin ~|~ \verb,\,\verb,<,latin\verb,>, ~|~ \verb,\,\verb,<,latin\,latin\verb,>, ~|~ greek ~|~ \\
       
    86            &   & \verb,\<^isub>, ~|~ \verb,\<^isup>, \\
       
    87     quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\
       
    88     latin & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\
       
    89     digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\
       
    90     sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~
       
    91      \verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \\
       
    92     & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~
       
    93     \verb,^, ~|~ \verb,_, ~|~ \verb,|, ~|~ \verb,~, \\
       
    94     greek & = & \verb,\<alpha>, ~|~ \verb,\<beta>, ~|~ \verb,\<gamma>, ~|~ \verb,\<delta>, ~| \\
       
    95           &   & \verb,\<epsilon>, ~|~ \verb,\<zeta>, ~|~ \verb,\<eta>, ~|~ \verb,\<theta>, ~| \\
       
    96           &   & \verb,\<iota>, ~|~ \verb,\<kappa>, ~|~ \verb,\<mu>, ~|~ \verb,\<nu>, ~| \\
       
    97           &   & \verb,\<xi>, ~|~ \verb,\<pi>, ~|~ \verb,\<rho>, ~|~ \verb,\<sigma>, ~|~ \verb,\<tau>, ~| \\
       
    98           &   & \verb,\<upsilon>, ~|~ \verb,\<phi>, ~|~ \verb,\<chi>, ~|~ \verb,\<psi>, ~| \\
       
    99           &   & \verb,\<omega>, ~|~ \verb,\<Gamma>, ~|~ \verb,\<Delta>, ~|~ \verb,\<Theta>, ~| \\
       
   100           &   & \verb,\<Lambda>, ~|~ \verb,\<Xi>, ~|~ \verb,\<Pi>, ~|~ \verb,\<Sigma>, ~| \\
       
   101           &   & \verb,\<Upsilon>, ~|~ \verb,\<Phi>, ~|~ \verb,\<Psi>, ~|~ \verb,\<Omega>, \\
       
   102   \end{matharray}
       
   103 
       
   104   The syntax of @{syntax string} admits any characters, including
       
   105   newlines; ``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash)
       
   106   need to be escaped by a backslash; arbitrary character codes may be
       
   107   specified as ``\verb|\|$ddd$'', with 3 decimal digits.  Alternative
       
   108   strings according to @{syntax altstring} are analogous, using single
       
   109   back-quotes instead.  The body of @{syntax verbatim} may consist of
       
   110   any text not containing ``\verb,*,\verb,},''; this allows convenient
       
   111   inclusion of quotes without further escapes.  The greek letters do
       
   112   \emph{not} include \verb,\<lambda>,, which is already used differently in
       
   113   the meta-logic.
       
   114 
       
   115   Common mathematical symbols such as @{text \<forall>} are represented in
       
   116   Isabelle as \verb,\<forall>,.  There are infinitely many Isabelle symbols
       
   117   like this, although proper presentation is left to front-end tools
       
   118   such as {\LaTeX} or Proof~General with the X-Symbol package.  A list
       
   119   of standard Isabelle symbols that work well with these tools is
       
   120   given in \cite[appendix~A]{isabelle-sys}.
       
   121   
       
   122   Source comments take the form \texttt{(*~\dots~*)} and may be
       
   123   nested, although user-interface tools might prevent this.  Note that
       
   124   \texttt{(*~\dots~*)} indicate source comments only, which are
       
   125   stripped after lexical analysis of the input.  The Isar document
       
   126   syntax also provides formal comments that are considered as part of
       
   127   the text (see \S\ref{sec:comments}).
       
   128 *}
       
   129 
       
   130 
       
   131 section {* Common syntax entities *}
       
   132 
       
   133 text {*
       
   134   We now introduce several basic syntactic entities, such as names,
       
   135   terms, and theorem specifications, which are factored out of the
       
   136   actual Isar language elements to be described later.
       
   137 *}
       
   138 
       
   139 
       
   140 subsection {* Names *}
       
   141 
       
   142 text {*
       
   143   Entity \railqtok{name} usually refers to any name of types,
       
   144   constants, theorems etc.\ that are to be \emph{declared} or
       
   145   \emph{defined} (so qualified identifiers are excluded here).  Quoted
       
   146   strings provide an escape for non-identifier names or those ruled
       
   147   out by outer syntax keywords (e.g.\ \verb|"let"|).  Already existing
       
   148   objects are usually referenced by \railqtok{nameref}.
       
   149 
       
   150   \indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
       
   151   \indexoutertoken{int}
       
   152   \begin{rail}
       
   153     name: ident | symident | string | nat
       
   154     ;
       
   155     parname: '(' name ')'
       
   156     ;
       
   157     nameref: name | longident
       
   158     ;
       
   159     int: nat | '-' nat
       
   160     ;
       
   161   \end{rail}
       
   162 *}
       
   163 
       
   164 
       
   165 subsection {* Comments \label{sec:comments} *}
       
   166 
       
   167 text {*
       
   168   Large chunks of plain \railqtok{text} are usually given
       
   169   \railtok{verbatim}, i.e.\ enclosed in
       
   170   \verb,{,\verb,*,~\dots~\verb,*,\verb,},.  For convenience, any of
       
   171   the smaller text units conforming to \railqtok{nameref} are admitted
       
   172   as well.  A marginal \railnonterm{comment} is of the form
       
   173   \texttt{--} \railqtok{text}.  Any number of these may occur within
       
   174   Isabelle/Isar commands.
       
   175 
       
   176   \indexoutertoken{text}\indexouternonterm{comment}
       
   177   \begin{rail}
       
   178     text: verbatim | nameref
       
   179     ;
       
   180     comment: '--' text
       
   181     ;
       
   182   \end{rail}
       
   183 *}
       
   184 
       
   185 
       
   186 subsection {* Type classes, sorts and arities *}
       
   187 
       
   188 text {*
       
   189   Classes are specified by plain names.  Sorts have a very simple
       
   190   inner syntax, which is either a single class name @{text c} or a
       
   191   list @{text "{c\<^sub>1, \<dots>, c\<^sub>n}"} referring to the
       
   192   intersection of these classes.  The syntax of type arities is given
       
   193   directly at the outer level.
       
   194 
       
   195   \railalias{subseteq}{\isasymsubseteq}
       
   196   \railterm{subseteq}
       
   197 
       
   198   \indexouternonterm{sort}\indexouternonterm{arity}
       
   199   \indexouternonterm{classdecl}
       
   200   \begin{rail}
       
   201     classdecl: name (('<' | subseteq) (nameref + ','))?
       
   202     ;
       
   203     sort: nameref
       
   204     ;
       
   205     arity: ('(' (sort + ',') ')')? sort
       
   206     ;
       
   207   \end{rail}
       
   208 *}
       
   209 
       
   210 
       
   211 subsection {* Types and terms \label{sec:types-terms} *}
       
   212 
       
   213 text {*
       
   214   The actual inner Isabelle syntax, that of types and terms of the
       
   215   logic, is far too sophisticated in order to be modelled explicitly
       
   216   at the outer theory level.  Basically, any such entity has to be
       
   217   quoted to turn it into a single token (the parsing and type-checking
       
   218   is performed internally later).  For convenience, a slightly more
       
   219   liberal convention is adopted: quotes may be omitted for any type or
       
   220   term that is already atomic at the outer level.  For example, one
       
   221   may just write \texttt{x} instead of \texttt{"x"}.  Note that
       
   222   symbolic identifiers (e.g.\ \texttt{++} or @{text "\<forall>"} are available
       
   223   as well, provided these have not been superseded by commands or
       
   224   other keywords already (e.g.\ \texttt{=} or \texttt{+}).
       
   225 
       
   226   \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
       
   227   \begin{rail}
       
   228     type: nameref | typefree | typevar
       
   229     ;
       
   230     term: nameref | var
       
   231     ;
       
   232     prop: term
       
   233     ;
       
   234   \end{rail}
       
   235 
       
   236   Positional instantiations are indicated by giving a sequence of
       
   237   terms, or the placeholder ``$\_$'' (underscore), which means to skip
       
   238   a position.
       
   239 
       
   240   \indexoutertoken{inst}\indexoutertoken{insts}
       
   241   \begin{rail}
       
   242     inst: underscore | term
       
   243     ;
       
   244     insts: (inst *)
       
   245     ;
       
   246   \end{rail}
       
   247 
       
   248   Type declarations and definitions usually refer to
       
   249   \railnonterm{typespec} on the left-hand side.  This models basic
       
   250   type constructor application at the outer syntax level.  Note that
       
   251   only plain postfix notation is available here, but no infixes.
       
   252 
       
   253   \indexouternonterm{typespec}
       
   254   \begin{rail}
       
   255     typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
       
   256     ;
       
   257   \end{rail}
       
   258 *}
       
   259 
       
   260 
       
   261 subsection {* Mixfix annotations *}
       
   262 
       
   263 text {*
       
   264   Mixfix annotations specify concrete \emph{inner} syntax of Isabelle
       
   265   types and terms.  Some commands such as @{command "types"} (see
       
   266   \S\ref{sec:types-pure}) admit infixes only, while @{command
       
   267   "consts"} (see \S\ref{sec:consts}) and @{command "syntax"} (see
       
   268   \S\ref{sec:syn-trans}) support the full range of general mixfixes
       
   269   and binders.
       
   270 
       
   271   \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
       
   272   \begin{rail}
       
   273     infix: '(' ('infix' | 'infixl' | 'infixr') string? nat ')'
       
   274     ;
       
   275     mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
       
   276     ;
       
   277     structmixfix: mixfix | '(' 'structure' ')'
       
   278     ;
       
   279 
       
   280     prios: '[' (nat + ',') ']'
       
   281     ;
       
   282   \end{rail}
       
   283 
       
   284   Here the \railtok{string} specifications refer to the actual mixfix
       
   285   template (see also \cite{isabelle-ref}), which may include literal
       
   286   text, spacing, blocks, and arguments (denoted by ``$_$''); the
       
   287   special symbol \verb,\<index>, (printed as ``\i'') represents an index
       
   288   argument that specifies an implicit structure reference (see also
       
   289   \S\ref{sec:locale}).  Infix and binder declarations provide common
       
   290   abbreviations for particular mixfix declarations.  So in practice,
       
   291   mixfix templates mostly degenerate to literal text for concrete
       
   292   syntax, such as ``\verb,++,'' for an infix symbol, or
       
   293   ``\verb,++,\i'' for an infix of an implicit structure.
       
   294 *}
       
   295 
       
   296 
       
   297 subsection {* Proof methods \label{sec:syn-meth} *}
       
   298 
       
   299 text {*
       
   300   Proof methods are either basic ones, or expressions composed of
       
   301   methods via ``\texttt{,}'' (sequential composition), ``\texttt{|}''
       
   302   (alternative choices), ``\texttt{?}'' (try), ``\texttt{+}'' (repeat
       
   303   at least once), ``\texttt{[$n$]}'' (restriction to first @{text n}
       
   304   sub-goals, default $n = 1$).  In practice, proof methods are usually
       
   305   just a comma separated list of \railqtok{nameref}~\railnonterm{args}
       
   306   specifications.  Note that parentheses may be dropped for single
       
   307   method specifications (with no arguments).
       
   308 
       
   309   \indexouternonterm{method}
       
   310   \begin{rail}
       
   311     method: (nameref | '(' methods ')') (() | '?' | '+' | '[' nat? ']')
       
   312     ;
       
   313     methods: (nameref args | method) + (',' | '|')
       
   314     ;
       
   315   \end{rail}
       
   316 
       
   317   Proper Isar proof methods do \emph{not} admit arbitrary goal
       
   318   addressing, but refer either to the first sub-goal or all sub-goals
       
   319   uniformly.  The goal restriction operator ``\texttt{[$n$]}''
       
   320   evaluates a method expression within a sandbox consisting of the
       
   321   first @{text n} sub-goals (which need to exist).  For example,
       
   322   @{text "simp_all[3]"} simplifies the first three sub-goals, while
       
   323   @{text "(rule foo, simp_all)[]"} simplifies all new goals that
       
   324   emerge from applying rule @{text "foo"} to the originally first one.
       
   325 
       
   326   Improper methods, notably tactic emulations, offer a separate
       
   327   low-level goal addressing scheme as explicit argument to the
       
   328   individual tactic being involved.  Here @{text "[!]"} refers to all
       
   329   goals, and @{text "[n-]"} to all goals starting from @{text "n"},
       
   330 
       
   331   \indexouternonterm{goalspec}
       
   332   \begin{rail}
       
   333     goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']'
       
   334     ;
       
   335   \end{rail}
       
   336 *}
       
   337 
       
   338 
       
   339 subsection {* Attributes and theorems \label{sec:syn-att} *}
       
   340 
       
   341 text {*
       
   342   Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their
       
   343   own ``semi-inner'' syntax, in the sense that input conforming to
       
   344   \railnonterm{args} below is parsed by the attribute a second time.
       
   345   The attribute argument specifications may be any sequence of atomic
       
   346   entities (identifiers, strings etc.), or properly bracketed argument
       
   347   lists.  Below \railqtok{atom} refers to any atomic entity, including
       
   348   any \railtok{keyword} conforming to \railtok{symident}.
       
   349 
       
   350   \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
       
   351   \begin{rail}
       
   352     atom: nameref | typefree | typevar | var | nat | keyword
       
   353     ;
       
   354     arg: atom | '(' args ')' | '[' args ']'
       
   355     ;
       
   356     args: arg *
       
   357     ;
       
   358     attributes: '[' (nameref args * ',') ']'
       
   359     ;
       
   360   \end{rail}
       
   361 
       
   362   Theorem specifications come in several flavors:
       
   363   \railnonterm{axmdecl} and \railnonterm{thmdecl} usually refer to
       
   364   axioms, assumptions or results of goal statements, while
       
   365   \railnonterm{thmdef} collects lists of existing theorems.  Existing
       
   366   theorems are given by \railnonterm{thmref} and
       
   367   \railnonterm{thmrefs}, the former requires an actual singleton
       
   368   result.
       
   369 
       
   370   There are three forms of theorem references:
       
   371   \begin{enumerate}
       
   372   
       
   373   \item named facts @{text "a"}
       
   374 
       
   375   \item selections from named facts @{text "a(i, j - k)"}
       
   376 
       
   377   \item literal fact propositions using @{syntax_ref altstring} syntax
       
   378   $\backquote\phi\backquote$, (see also method @{method_ref fact} in
       
   379   \S\ref{sec:pure-meth-att}).
       
   380 
       
   381   \end{enumerate}
       
   382 
       
   383   Any kind of theorem specification may include lists of attributes
       
   384   both on the left and right hand sides; attributes are applied to any
       
   385   immediately preceding fact.  If names are omitted, the theorems are
       
   386   not stored within the theorem database of the theory or proof
       
   387   context, but any given attributes are applied nonetheless.
       
   388 
       
   389   An extra pair of brackets around attribute declarations --- such as
       
   390   ``@{text "[[simproc a]]"}'' --- abbreviates a theorem reference
       
   391   involving an internal dummy fact, which will be ignored later on.
       
   392   So only the effect of the attribute on the background context will
       
   393   persist.  This form of in-place declarations is particularly useful
       
   394   with commands like @{command "declare"} and @{command "using"}.
       
   395 
       
   396   \indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
       
   397   \indexouternonterm{thmdef}\indexouternonterm{thmref}
       
   398   \indexouternonterm{thmrefs}\indexouternonterm{selection}
       
   399   \begin{rail}
       
   400     axmdecl: name attributes? ':'
       
   401     ;
       
   402     thmdecl: thmbind ':'
       
   403     ;
       
   404     thmdef: thmbind '='
       
   405     ;
       
   406     thmref: (nameref selection? | altstring) attributes? | '[' attributes ']'
       
   407     ;
       
   408     thmrefs: thmref +
       
   409     ;
       
   410 
       
   411     thmbind: name attributes | name | attributes
       
   412     ;
       
   413     selection: '(' ((nat | nat '-' nat?) + ',') ')'
       
   414     ;
       
   415   \end{rail}
       
   416 *}
       
   417 
       
   418 
       
   419 subsection {* Term patterns and declarations \label{sec:term-decls} *}
       
   420 
       
   421 text {*
       
   422   Wherever explicit propositions (or term fragments) occur in a proof
       
   423   text, casual binding of schematic term variables may be given
       
   424   specified via patterns of the form ``@{text "(\<IS> p\<^sub>1 \<dots>
       
   425   p\<^sub>n)"}''.  This works both for \railqtok{term} and \railqtok{prop}.
       
   426 
       
   427   \indexouternonterm{termpat}\indexouternonterm{proppat}
       
   428   \begin{rail}
       
   429     termpat: '(' ('is' term +) ')'
       
   430     ;
       
   431     proppat: '(' ('is' prop +) ')'
       
   432     ;
       
   433   \end{rail}
       
   434 
       
   435   \medskip Declarations of local variables @{text "x :: \<tau>"} and
       
   436   logical propositions @{text "a : \<phi>"} represent different views on
       
   437   the same principle of introducing a local scope.  In practice, one
       
   438   may usually omit the typing of \railnonterm{vars} (due to
       
   439   type-inference), and the naming of propositions (due to implicit
       
   440   references of current facts).  In any case, Isar proof elements
       
   441   usually admit to introduce multiple such items simultaneously.
       
   442 
       
   443   \indexouternonterm{vars}\indexouternonterm{props}
       
   444   \begin{rail}
       
   445     vars: (name+) ('::' type)?
       
   446     ;
       
   447     props: thmdecl? (prop proppat? +)
       
   448     ;
       
   449   \end{rail}
       
   450 
       
   451   The treatment of multiple declarations corresponds to the
       
   452   complementary focus of \railnonterm{vars} versus
       
   453   \railnonterm{props}.  In ``@{text "x\<^sub>1 \<dots> x\<^sub>n :: \<tau>"}''
       
   454   the typing refers to all variables, while in @{text "a: \<phi>\<^sub>1 \<dots>
       
   455   \<phi>\<^sub>n"} the naming refers to all propositions collectively.
       
   456   Isar language elements that refer to \railnonterm{vars} or
       
   457   \railnonterm{props} typically admit separate typings or namings via
       
   458   another level of iteration, with explicit @{keyword_ref "and"}
       
   459   separators; e.g.\ see @{command "fix"} and @{command "assume"} in
       
   460   \S\ref{sec:proof-context}.
       
   461 *}
       
   462 
       
   463 
       
   464 subsection {* Antiquotations \label{sec:antiq} *}
       
   465 
       
   466 text {*
       
   467   \begin{matharray}{rcl}
       
   468     @{antiquotation_def "theory"} & : & \isarantiq \\
       
   469     @{antiquotation_def "thm"} & : & \isarantiq \\
       
   470     @{antiquotation_def "prop"} & : & \isarantiq \\
       
   471     @{antiquotation_def "term"} & : & \isarantiq \\
       
   472     @{antiquotation_def const} & : & \isarantiq \\
       
   473     @{antiquotation_def abbrev} & : & \isarantiq \\
       
   474     @{antiquotation_def typeof} & : & \isarantiq \\
       
   475     @{antiquotation_def typ} & : & \isarantiq \\
       
   476     @{antiquotation_def thm_style} & : & \isarantiq \\
       
   477     @{antiquotation_def term_style} & : & \isarantiq \\
       
   478     @{antiquotation_def "text"} & : & \isarantiq \\
       
   479     @{antiquotation_def goals} & : & \isarantiq \\
       
   480     @{antiquotation_def subgoals} & : & \isarantiq \\
       
   481     @{antiquotation_def prf} & : & \isarantiq \\
       
   482     @{antiquotation_def full_prf} & : & \isarantiq \\
       
   483     @{antiquotation_def ML} & : & \isarantiq \\
       
   484     @{antiquotation_def ML_type} & : & \isarantiq \\
       
   485     @{antiquotation_def ML_struct} & : & \isarantiq \\
       
   486   \end{matharray}
       
   487 
       
   488   The text body of formal comments (see also \S\ref{sec:comments}) may
       
   489   contain antiquotations of logical entities, such as theorems, terms
       
   490   and types, which are to be presented in the final output produced by
       
   491   the Isabelle document preparation system (see also
       
   492   \S\ref{sec:document-prep}).
       
   493 
       
   494   Thus embedding of ``@{text "@{term [show_types] \"f x = a + x\"}"}''
       
   495   within a text block would cause
       
   496   \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} to appear in the final {\LaTeX} document.  Also note that theorem
       
   497   antiquotations may involve attributes as well.  For example,
       
   498   \texttt{{\at}{\ttlbrace}thm~sym~[no_vars]{\ttrbrace}} would print
       
   499   the statement where all schematic variables have been replaced by
       
   500   fixed ones, which are easier to read.
       
   501 
       
   502   \begin{rail}
       
   503     atsign lbrace antiquotation rbrace
       
   504     ;
       
   505 
       
   506     antiquotation:
       
   507       'theory' options name |
       
   508       'thm' options thmrefs |
       
   509       'prop' options prop |
       
   510       'term' options term |
       
   511       'const' options term |
       
   512       'abbrev' options term |
       
   513       'typeof' options term |
       
   514       'typ' options type |
       
   515       'thm\_style' options name thmref |
       
   516       'term\_style' options name term |
       
   517       'text' options name |
       
   518       'goals' options |
       
   519       'subgoals' options |
       
   520       'prf' options thmrefs |
       
   521       'full\_prf' options thmrefs |
       
   522       'ML' options name |
       
   523       'ML\_type' options name |
       
   524       'ML\_struct' options name
       
   525     ;
       
   526     options: '[' (option * ',') ']'
       
   527     ;
       
   528     option: name | name '=' name
       
   529     ;
       
   530   \end{rail}
       
   531 
       
   532   Note that the syntax of antiquotations may \emph{not} include source
       
   533   comments \texttt{(*~\dots~*)} or verbatim text
       
   534   \verb|{*|~\dots~\verb|*|\verb|}|.
       
   535 
       
   536   \begin{descr}
       
   537   
       
   538   \item [@{text "@{theory A}"}] prints the name @{text "A"}, which is
       
   539   guaranteed to refer to a valid ancestor theory in the current
       
   540   context.
       
   541 
       
   542   \item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems @{text
       
   543   "a\<^sub>1 \<dots> a\<^sub>n"}.  Note that attribute specifications may be
       
   544   included as well (see also \S\ref{sec:syn-att}); the @{attribute_ref
       
   545   no_vars} rule (see \S\ref{sec:misc-meth-att}) would be particularly
       
   546   useful to suppress printing of schematic variables.
       
   547 
       
   548   \item [@{text "@{prop \<phi>}"}] prints a well-typed proposition @{text
       
   549   "\<phi>"}.
       
   550 
       
   551   \item [@{text "@{term t}"}] prints a well-typed term @{text "t"}.
       
   552 
       
   553   \item [@{text "@{const c}"}] prints a logical or syntactic constant
       
   554   @{text "c"}.
       
   555   
       
   556   \item [@{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"}] prints a constant
       
   557   abbreviation @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in
       
   558   the current context.
       
   559 
       
   560   \item [@{text "@{typeof t}"}] prints the type of a well-typed term
       
   561   @{text "t"}.
       
   562 
       
   563   \item [@{text "@{typ \<tau>}"}] prints a well-formed type @{text "\<tau>"}.
       
   564   
       
   565   \item [@{text "@{thm_style s a}"}] prints theorem @{text a},
       
   566   previously applying a style @{text s} to it (see below).
       
   567   
       
   568   \item [@{text "@{term_style s t}"}] prints a well-typed term @{text
       
   569   t} after applying a style @{text s} to it (see below).
       
   570 
       
   571   \item [@{text "@{text s}"}] prints uninterpreted source text @{text
       
   572   s}.  This is particularly useful to print portions of text according
       
   573   to the Isabelle {\LaTeX} output style, without demanding
       
   574   well-formedness (e.g.\ small pieces of terms that should not be
       
   575   parsed or type-checked yet).
       
   576 
       
   577   \item [@{text "@{goals}"}] prints the current \emph{dynamic} goal
       
   578   state.  This is mainly for support of tactic-emulation scripts
       
   579   within Isar --- presentation of goal states does not conform to
       
   580   actual human-readable proof documents.
       
   581 
       
   582   Please do not include goal states into document output unless you
       
   583   really know what you are doing!
       
   584   
       
   585   \item [@{text "@{subgoals}"}] is similar to @{text "@{goals}"}, but
       
   586   does not print the main goal.
       
   587   
       
   588   \item [@{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}] prints the (compact)
       
   589   proof terms corresponding to the theorems @{text "a\<^sub>1 \<dots>
       
   590   a\<^sub>n"}. Note that this requires proof terms to be switched on
       
   591   for the current object logic (see the ``Proof terms'' section of the
       
   592   Isabelle reference manual for information on how to do this).
       
   593   
       
   594   \item [@{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"}] is like @{text
       
   595   "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}, but displays the full proof terms,
       
   596   i.e.\ also displays information omitted in the compact proof term,
       
   597   which is denoted by ``$_$'' placeholders there.
       
   598   
       
   599   \item [@{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
       
   600   "@{ML_struct s}"}] check text @{text s} as ML value, type, and
       
   601   structure, respectively.  The source is displayed verbatim.
       
   602 
       
   603   \end{descr}
       
   604 
       
   605   \medskip The following standard styles for use with @{text
       
   606   thm_style} and @{text term_style} are available:
       
   607 
       
   608   \begin{descr}
       
   609   
       
   610   \item [@{text lhs}] extracts the first argument of any application
       
   611   form with at least two arguments -- typically meta-level or
       
   612   object-level equality, or any other binary relation.
       
   613   
       
   614   \item [@{text rhs}] is like @{text lhs}, but extracts the second
       
   615   argument.
       
   616   
       
   617   \item [@{text "concl"}] extracts the conclusion @{text C} from a rule
       
   618   in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
       
   619   
       
   620   \item [@{text "prem1"}, \dots, @{text "prem9"}] extract premise
       
   621   number $1$, \dots, $9$, respectively, from from a rule in
       
   622   Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
       
   623 
       
   624   \end{descr}
       
   625 
       
   626   \medskip
       
   627   The following options are available to tune the output.  Note that most of
       
   628   these coincide with ML flags of the same names (see also \cite{isabelle-ref}).
       
   629 
       
   630   \begin{descr}
       
   631 
       
   632   \item[@{text "show_types = bool"} and @{text "show_sorts = bool"}]
       
   633   control printing of explicit type and sort constraints.
       
   634 
       
   635   \item[@{text "show_structs = bool"}] controls printing of implicit
       
   636   structures.
       
   637 
       
   638   \item[@{text "long_names = bool"}] forces names of types and
       
   639   constants etc.\ to be printed in their fully qualified internal
       
   640   form.
       
   641 
       
   642   \item[@{text "short_names = bool"}] forces names of types and
       
   643   constants etc.\ to be printed unqualified.  Note that internalizing
       
   644   the output again in the current context may well yield a different
       
   645   result.
       
   646 
       
   647   \item[@{text "unique_names = bool"}] determines whether the printed
       
   648   version of qualified names should be made sufficiently long to avoid
       
   649   overlap with names declared further back.  Set to @{text false} for
       
   650   more concise output.
       
   651 
       
   652   \item[@{text "eta_contract = bool"}] prints terms in @{text
       
   653   \<eta>}-contracted form.
       
   654 
       
   655   \item[@{text "display = bool"}] indicates if the text is to be
       
   656   output as multi-line ``display material'', rather than a small piece
       
   657   of text without line breaks (which is the default).
       
   658 
       
   659   \item[@{text "break = bool"}] controls line breaks in non-display
       
   660   material.
       
   661 
       
   662   \item[@{text "quotes = bool"}] indicates if the output should be
       
   663   enclosed in double quotes.
       
   664 
       
   665   \item[@{text "mode = name"}] adds @{text name} to the print mode to
       
   666   be used for presentation (see also \cite{isabelle-ref}).  Note that
       
   667   the standard setup for {\LaTeX} output is already present by
       
   668   default, including the modes @{text latex} and @{text xsymbols}.
       
   669 
       
   670   \item[@{text "margin = nat"} and @{text "indent = nat"}] change the
       
   671   margin or indentation for pretty printing of display material.
       
   672 
       
   673   \item[@{text "source = bool"}] prints the source text of the
       
   674   antiquotation arguments, rather than the actual value.  Note that
       
   675   this does not affect well-formedness checks of @{antiquotation
       
   676   "thm"}, @{antiquotation "term"}, etc. (only the @{antiquotation
       
   677   "text"} antiquotation admits arbitrary output).
       
   678 
       
   679   \item[@{text "goals_limit = nat"}] determines the maximum number of
       
   680   goals to be printed.
       
   681 
       
   682   \item[@{text "locale = name"}] specifies an alternative locale
       
   683   context used for evaluating and printing the subsequent argument.
       
   684 
       
   685   \end{descr}
       
   686 
       
   687   For boolean flags, ``@{text "name = true"}'' may be abbreviated as
       
   688   ``@{text name}''.  All of the above flags are disabled by default,
       
   689   unless changed from ML.
       
   690 
       
   691   \medskip Note that antiquotations do not only spare the author from
       
   692   tedious typing of logical entities, but also achieve some degree of
       
   693   consistency-checking of informal explanations with formal
       
   694   developments: well-formedness of terms and types with respect to the
       
   695   current theory or proof context is ensured here.
       
   696 *}
       
   697 
       
   698 
       
   699 subsection {* Tagged commands \label{sec:tags} *}
       
   700 
       
   701 text {*
       
   702   Each Isabelle/Isar command may be decorated by presentation tags:
       
   703 
       
   704   \indexouternonterm{tags}
       
   705   \begin{rail}
       
   706     tags: ( tag * )
       
   707     ;
       
   708     tag: '\%' (ident | string)
       
   709   \end{rail}
       
   710 
       
   711   The tags @{text "theory"}, @{text "proof"}, @{text "ML"} are already
       
   712   pre-declared for certain classes of commands:
       
   713 
       
   714  \medskip
       
   715 
       
   716   \begin{tabular}{ll}
       
   717     @{text "theory"} & theory begin/end \\
       
   718     @{text "proof"} & all proof commands \\
       
   719     @{text "ML"} & all commands involving ML code \\
       
   720   \end{tabular}
       
   721 
       
   722   \medskip The Isabelle document preparation system (see also
       
   723   \cite{isabelle-sys}) allows tagged command regions to be presented
       
   724   specifically, e.g.\ to fold proof texts, or drop parts of the text
       
   725   completely.
       
   726 
       
   727   For example ``@{command "by"}~@{text "%invisible auto"}'' would
       
   728   cause that piece of proof to be treated as @{text invisible} instead
       
   729   of @{text "proof"} (the default), which may be either show or hidden
       
   730   depending on the document setup.  In contrast, ``@{command
       
   731   "by"}~@{text "%visible auto"}'' would force this text to be shown
       
   732   invariably.
       
   733 
       
   734   Explicit tag specifications within a proof apply to all subsequent
       
   735   commands of the same level of nesting.  For example, ``@{command
       
   736   "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' would force the
       
   737   whole sub-proof to be typeset as @{text visible} (unless some of its
       
   738   parts are tagged differently).
       
   739 *}
       
   740 
       
   741 end