doc-src/IsarRef/Thy/Document_Preparation.thy
author wenzelm
Mon, 29 Nov 2010 11:22:40 +0100
changeset 40801 6cfacec435e6
parent 40800 330eb65c9469
child 40879 ca132ef44944
permissions -rw-r--r--
added document antiquotation @{file};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     1
theory Document_Preparation
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     2
imports Main
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     3
begin
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     4
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     5
chapter {* Document preparation \label{ch:document-prep} *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     6
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
     7
text {* Isabelle/Isar provides a simple document preparation system
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
     8
  based on regular {PDF-\LaTeX} technology, with full support for
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
     9
  hyper-links and bookmarks.  Thus the results are well suited for WWW
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    10
  browsing and as printed copies.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    11
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    12
  \medskip Isabelle generates {\LaTeX} output while running a
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    13
  \emph{logic session} (see also \cite{isabelle-sys}).  Getting
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    14
  started with a working configuration for common situations is quite
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    15
  easy by using the Isabelle @{verbatim mkdir} and @{verbatim make}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    16
  tools.  First invoke
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    17
\begin{ttbox}
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 27881
diff changeset
    18
  isabelle mkdir Foo
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    19
\end{ttbox}
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    20
  to initialize a separate directory for session @{verbatim Foo} (it
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    21
  is safe to experiment, since @{verbatim "isabelle mkdir"} never
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    22
  overwrites existing files).  Ensure that @{verbatim "Foo/ROOT.ML"}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    23
  holds ML commands to load all theories required for this session;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    24
  furthermore @{verbatim "Foo/document/root.tex"} should include any
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    25
  special {\LaTeX} macro packages required for your document (the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    26
  default is usually sufficient as a start).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    27
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    28
  The session is controlled by a separate @{verbatim IsaMakefile}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    29
  (with crude source dependencies by default).  This file is located
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    30
  one level up from the @{verbatim Foo} directory location.  Now
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    31
  invoke
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    32
\begin{ttbox}
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 27881
diff changeset
    33
  isabelle make Foo
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    34
\end{ttbox}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    35
  to run the @{verbatim Foo} session, with browser information and
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    36
  document preparation enabled.  Unless any errors are reported by
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    37
  Isabelle or {\LaTeX}, the output will appear inside the directory
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    38
  defined by the @{verbatim ISABELLE_BROWSER_INFO} setting (as
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    39
  reported by the batch job in verbose mode).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    40
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    41
  \medskip You may also consider to tune the @{verbatim usedir}
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    42
  options in @{verbatim IsaMakefile}, for example to switch the output
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    43
  format between @{verbatim pdf} and @{verbatim dvi}, or activate the
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    44
  @{verbatim "-D"} option to retain a second copy of the generated
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    45
  {\LaTeX} sources (for manual inspection or separate runs of
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    46
  @{executable latex}).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    47
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    48
  \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    49
  for further details on Isabelle logic sessions and theory
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    50
  presentation.  The Isabelle/HOL tutorial \cite{isabelle-hol-book}
28746
c1b151a60a66 tuned intro of "Document preparation";
wenzelm
parents: 28504
diff changeset
    51
  also covers theory presentation to some extent.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    52
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    53
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    54
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    55
section {* Markup commands \label{sec:markup} *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    56
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    57
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    58
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    59
    @{command_def "header"} & : & @{text "toplevel \<rightarrow> toplevel"} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    60
    @{command_def "chapter"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    61
    @{command_def "section"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    62
    @{command_def "subsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    63
    @{command_def "subsubsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    64
    @{command_def "text"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    65
    @{command_def "text_raw"} & : & @{text "local_theory \<rightarrow> local_theory"} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    66
    @{command_def "sect"} & : & @{text "proof \<rightarrow> proof"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    67
    @{command_def "subsect"} & : & @{text "proof \<rightarrow> proof"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    68
    @{command_def "subsubsect"} & : & @{text "proof \<rightarrow> proof"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    69
    @{command_def "txt"} & : & @{text "proof \<rightarrow> proof"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    70
    @{command_def "txt_raw"} & : & @{text "proof \<rightarrow> proof"} \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    71
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    72
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    73
  Markup commands provide a structured way to insert text into the
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    74
  document generated from a theory.  Each markup command takes a
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    75
  single @{syntax text} argument, which is passed as argument to a
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    76
  corresponding {\LaTeX} macro.  The default macros provided by
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40255
diff changeset
    77
  @{file "~~/lib/texinputs/isabelle.sty"} can be redefined according
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    78
  to the needs of the underlying document and {\LaTeX} styles.
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    79
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    80
  Note that formal comments (\secref{sec:comments}) are similar to
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    81
  markup commands, but have a different status within Isabelle/Isar
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    82
  syntax.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    83
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    84
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    85
    ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    86
    ;
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 39689
diff changeset
    87
    ('header' | 'text_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt_raw') text
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    88
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    89
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    90
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
    91
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    92
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
    93
  \item @{command header} provides plain text markup just preceding
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    94
  the formal beginning of a theory.  The corresponding {\LaTeX} macro
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    95
  is @{verbatim "\\isamarkupheader"}, which acts like @{command
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    96
  section} by default.
27049
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    97
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
    98
  \item @{command chapter}, @{command section}, @{command subsection},
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
    99
  and @{command subsubsection} mark chapter and section headings
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   100
  within the main theory body or local theory targets.  The
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   101
  corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"},
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   102
  @{verbatim "\\isamarkupsection"}, @{verbatim
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   103
  "\\isamarkupsubsection"} etc.
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   104
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   105
  \item @{command sect}, @{command subsect}, and @{command subsubsect}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   106
  mark section headings within proofs.  The corresponding {\LaTeX}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   107
  macros are @{verbatim "\\isamarkupsect"}, @{verbatim
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   108
  "\\isamarkupsubsect"} etc.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   109
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   110
  \item @{command text} and @{command txt} specify paragraphs of plain
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   111
  text.  This corresponds to a {\LaTeX} environment @{verbatim
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   112
  "\\begin{isamarkuptext}"} @{text "\<dots>"} @{verbatim
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   113
  "\\end{isamarkuptext}"} etc.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   114
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   115
  \item @{command text_raw} and @{command txt_raw} insert {\LaTeX}
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   116
  source into the output, without additional markup.  Thus the full
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   117
  range of document manipulations becomes available, at the risk of
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   118
  messing up document output.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   119
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   120
  \end{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   121
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   122
  Except for @{command "text_raw"} and @{command "txt_raw"}, the text
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
   123
  passed to any of the above markup commands may refer to formal
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   124
  entities via \emph{document antiquotations}, see also
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   125
  \secref{sec:antiq}.  These are interpreted in the present theory or
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   126
  proof context, or the named @{text "target"}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   127
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   128
  \medskip The proof markup commands closely resemble those for theory
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   129
  specifications, but have a different formal status and produce
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   130
  different {\LaTeX} macros.  The default definitions coincide for
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   131
  analogous commands such as @{command section} and @{command sect}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   132
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   133
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   134
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   135
section {* Document Antiquotations \label{sec:antiq} *}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   136
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   137
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   138
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   139
    @{antiquotation_def "theory"} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   140
    @{antiquotation_def "thm"} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   141
    @{antiquotation_def "lemma"} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   142
    @{antiquotation_def "prop"} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   143
    @{antiquotation_def "term"} & : & @{text antiquotation} \\
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   144
    @{antiquotation_def term_type} & : & @{text antiquotation} \\
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   145
    @{antiquotation_def typeof} & : & @{text antiquotation} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   146
    @{antiquotation_def const} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   147
    @{antiquotation_def abbrev} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   148
    @{antiquotation_def typ} & : & @{text antiquotation} \\
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   149
    @{antiquotation_def type} & : & @{text antiquotation} \\
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   150
    @{antiquotation_def class} & : & @{text antiquotation} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   151
    @{antiquotation_def "text"} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   152
    @{antiquotation_def goals} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   153
    @{antiquotation_def subgoals} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   154
    @{antiquotation_def prf} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   155
    @{antiquotation_def full_prf} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   156
    @{antiquotation_def ML} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   157
    @{antiquotation_def ML_type} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   158
    @{antiquotation_def ML_struct} & : & @{text antiquotation} \\
40801
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   159
    @{antiquotation_def "file"} & : & @{text antiquotation} \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   160
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   161
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   162
  The overall content of an Isabelle/Isar theory may alternate between
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   163
  formal and informal text.  The main body consists of formal
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   164
  specification and proof commands, interspersed with markup commands
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   165
  (\secref{sec:markup}) or document comments (\secref{sec:comments}).
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   166
  The argument of markup commands quotes informal text to be printed
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   167
  in the resulting document, but may again refer to formal entities
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   168
  via \emph{document antiquotations}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   169
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   170
  For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}''
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   171
  within a text block makes
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   172
  \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} appear in the final {\LaTeX} document.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   173
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   174
  Antiquotations usually spare the author tedious typing of logical
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   175
  entities in full detail.  Even more importantly, some degree of
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   176
  consistency-checking between the main body of formal text and its
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   177
  informal explanation is achieved, since terms and types appearing in
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   178
  antiquotations are checked within the current theory or proof
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   179
  context.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   180
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   181
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   182
    atsign lbrace antiquotation rbrace
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   183
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   184
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   185
    antiquotation:
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   186
      'theory' options name |
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   187
      'thm' options styles thmrefs |
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27053
diff changeset
   188
      'lemma' options prop 'by' method |
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   189
      'prop' options styles prop |
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   190
      'term' options styles term |
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   191
      'term_type' options styles term |
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   192
      'typeof' options styles term |
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   193
      'const' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   194
      'abbrev' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   195
      'typ' options type |
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   196
      'type' options name |
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   197
      'class' options name |
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   198
      'text' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   199
      'goals' options |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   200
      'subgoals' options |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   201
      'prf' options thmrefs |
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 39689
diff changeset
   202
      'full_prf' options thmrefs |
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   203
      'ML' options name |
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 39689
diff changeset
   204
      'ML_type' options name |
40801
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   205
      'ML_struct' options name |
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   206
      'file' options name
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   207
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   208
    options: '[' (option * ',') ']'
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   209
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   210
    option: name | name '=' name
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   211
    ;
32892
9742f6130f10 corrected syntax diagrams for styles
haftmann
parents: 32891
diff changeset
   212
    styles: '(' (style + ',') ')'
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   213
    ;
32892
9742f6130f10 corrected syntax diagrams for styles
haftmann
parents: 32891
diff changeset
   214
    style: (name +)
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   215
    ;
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   216
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   217
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   218
  Note that the syntax of antiquotations may \emph{not} include source
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   219
  comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   220
  text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   221
  "*"}@{verbatim "}"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   222
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   223
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   224
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   225
  \item @{text "@{theory A}"} prints the name @{text "A"}, which is
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   226
  guaranteed to refer to a valid ancestor theory in the current
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   227
  context.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   228
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   229
  \item @{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   230
  Full fact expressions are allowed here, including attributes
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   231
  (\secref{sec:syn-att}).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   232
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   233
  \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   234
  "\<phi>"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   235
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   236
  \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   237
  @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27053
diff changeset
   238
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   239
  \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   240
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   241
  \item @{text "@{term_type t}"} prints a well-typed term @{text "t"}
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   242
  annotated with its type.
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   243
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   244
  \item @{text "@{typeof t}"} prints the type of a well-typed term
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   245
  @{text "t"}.
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   246
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   247
  \item @{text "@{const c}"} prints a logical or syntactic constant
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   248
  @{text "c"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   249
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   250
  \item @{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"} prints a constant abbreviation
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   251
  @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   252
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   253
  \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   254
39689
78b185bf7660 clarified @{type} antiquotation: abbreviations and nonterminals count as "syntactic", disallow TFrees;
wenzelm
parents: 39305
diff changeset
   255
  \item @{text "@{type \<kappa>}"} prints a (logical or syntactic) type
78b185bf7660 clarified @{type} antiquotation: abbreviations and nonterminals count as "syntactic", disallow TFrees;
wenzelm
parents: 39305
diff changeset
   256
    constructor @{text "\<kappa>"}.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   257
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   258
  \item @{text "@{class c}"} prints a class @{text c}.
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   259
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   260
  \item @{text "@{text s}"} prints uninterpreted source text @{text
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   261
  s}.  This is particularly useful to print portions of text according
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   262
  to the Isabelle document style, without demanding well-formedness,
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   263
  e.g.\ small pieces of terms that should not be parsed or
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   264
  type-checked yet.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   265
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   266
  \item @{text "@{goals}"} prints the current \emph{dynamic} goal
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   267
  state.  This is mainly for support of tactic-emulation scripts
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   268
  within Isar.  Presentation of goal states does not conform to the
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   269
  idea of human-readable proof documents!
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   270
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   271
  When explaining proofs in detail it is usually better to spell out
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   272
  the reasoning via proper Isar proof commands, instead of peeking at
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   273
  the internal machine configuration.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   274
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   275
  \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   276
  does not print the main goal.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   277
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   278
  \item @{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"} prints the (compact) proof terms
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   279
  corresponding to the theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that this
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   280
  requires proof terms to be switched on for the current logic
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   281
  session.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   282
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   283
  \item @{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \<dots>
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   284
  a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   285
  information omitted in the compact proof term, which is denoted by
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   286
  ``@{text _}'' placeholders there.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   287
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   288
  \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   289
  "@{ML_struct s}"} check text @{text s} as ML value, type, and
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   290
  structure, respectively.  The source is printed verbatim.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   291
40801
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   292
  \item @{text "@{file path}"} checks that @{text "path"} refers to a
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   293
  file (or directory) and prints it verbatim.
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   294
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   295
  \end{description}
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   296
*}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   297
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   298
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   299
subsubsection {* Styled antiquotations *}
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   300
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   301
text {* The antiquotations @{text thm}, @{text prop} and @{text
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   302
  term} admit an extra \emph{style} specification to modify the
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   303
  printed result.  A style is specified by a name with a possibly
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   304
  empty number of arguments;  multiple styles can be sequenced with
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   305
  commas.  The following standard styles are available:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   306
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   307
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   308
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   309
  \item @{text lhs} extracts the first argument of any application
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   310
  form with at least two arguments --- typically meta-level or
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   311
  object-level equality, or any other binary relation.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   312
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   313
  \item @{text rhs} is like @{text lhs}, but extracts the second
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   314
  argument.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   315
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   316
  \item @{text "concl"} extracts the conclusion @{text C} from a rule
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   317
  in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   318
  
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   319
  \item @{text "prem"} @{text n} extract premise number
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   320
  @{text "n"} from from a rule in Horn-clause
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   321
  normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   322
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   323
  \end{description}
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   324
*}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   325
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   326
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   327
subsubsection {* General options *}
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   328
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   329
text {* The following options are available to tune the printed output
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   330
  of antiquotations.  Note that many of these coincide with global ML
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   331
  flags of the same names.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   332
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   333
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   334
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   335
  \item @{antiquotation_option_def show_types}~@{text "= bool"} and
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   336
  @{antiquotation_option_def show_sorts}~@{text "= bool"} control
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   337
  printing of explicit type and sort constraints.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   338
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   339
  \item @{antiquotation_option_def show_structs}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   340
  controls printing of implicit structures.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   341
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   342
  \item @{antiquotation_option_def long_names}~@{text "= bool"} forces
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   343
  names of types and constants etc.\ to be printed in their fully
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   344
  qualified internal form.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   345
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   346
  \item @{antiquotation_option_def short_names}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   347
  forces names of types and constants etc.\ to be printed unqualified.
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   348
  Note that internalizing the output again in the current context may
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   349
  well yield a different result.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   350
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   351
  \item @{antiquotation_option_def unique_names}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   352
  determines whether the printed version of qualified names should be
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   353
  made sufficiently long to avoid overlap with names declared further
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   354
  back.  Set to @{text false} for more concise output.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   355
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   356
  \item @{antiquotation_option_def eta_contract}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   357
  prints terms in @{text \<eta>}-contracted form.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   358
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   359
  \item @{antiquotation_option_def display}~@{text "= bool"} indicates
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   360
  if the text is to be output as multi-line ``display material'',
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   361
  rather than a small piece of text without line breaks (which is the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   362
  default).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   363
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   364
  In this mode the embedded entities are printed in the same style as
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   365
  the main theory text.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   366
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   367
  \item @{antiquotation_option_def break}~@{text "= bool"} controls
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   368
  line breaks in non-display material.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   369
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   370
  \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   371
  if the output should be enclosed in double quotes.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   372
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   373
  \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   374
  name} to the print mode to be used for presentation.  Note that the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   375
  standard setup for {\LaTeX} output is already present by default,
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   376
  including the modes @{text latex} and @{text xsymbols}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   377
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   378
  \item @{antiquotation_option_def margin}~@{text "= nat"} and
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   379
  @{antiquotation_option_def indent}~@{text "= nat"} change the margin
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   380
  or indentation for pretty printing of display material.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   381
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   382
  \item @{antiquotation_option_def goals_limit}~@{text "= nat"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   383
  determines the maximum number of goals to be printed (for goal-based
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   384
  antiquotation).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   385
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   386
  \item @{antiquotation_option_def source}~@{text "= bool"} prints the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   387
  original source text of the antiquotation arguments, rather than its
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   388
  internal representation.  Note that formal checking of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   389
  @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   390
  enabled; use the @{antiquotation "text"} antiquotation for unchecked
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   391
  output.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   392
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   393
  Regular @{text "term"} and @{text "typ"} antiquotations with @{text
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   394
  "source = false"} involve a full round-trip from the original source
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   395
  to an internalized logical entity back to a source form, according
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   396
  to the syntax of the current context.  Thus the printed output is
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   397
  not under direct control of the author, it may even fluctuate a bit
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   398
  as the underlying theory is changed later on.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   399
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   400
  In contrast, @{antiquotation_option_def source}~@{text "= true"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   401
  admits direct printing of the given source text, with the desirable
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   402
  well-formedness check in the background, but without modification of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   403
  the printed text.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   404
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   405
  \end{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   406
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   407
  For boolean flags, ``@{text "name = true"}'' may be abbreviated as
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   408
  ``@{text name}''.  All of the above flags are disabled by default,
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   409
  unless changed from ML, say in the @{verbatim "ROOT.ML"} of the
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   410
  logic session.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   411
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   412
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   413
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   414
section {* Markup via command tags \label{sec:tags} *}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   415
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   416
text {* Each Isabelle/Isar command may be decorated by additional
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   417
  presentation tags, to indicate some modification in the way it is
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   418
  printed in the document.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   419
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   420
  \indexouternonterm{tags}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   421
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   422
    tags: ( tag * )
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   423
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   424
    tag: '\%' (ident | string)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   425
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   426
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   427
  Some tags are pre-declared for certain classes of commands, serving
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   428
  as default markup if no tags are given in the text:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   429
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   430
  \medskip
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   431
  \begin{tabular}{ll}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   432
    @{text "theory"} & theory begin/end \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   433
    @{text "proof"} & all proof commands \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   434
    @{text "ML"} & all commands involving ML code \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   435
  \end{tabular}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   436
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   437
  \medskip The Isabelle document preparation system
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   438
  \cite{isabelle-sys} allows tagged command regions to be presented
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   439
  specifically, e.g.\ to fold proof texts, or drop parts of the text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   440
  completely.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   441
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   442
  For example ``@{command "by"}~@{text "%invisible auto"}'' causes
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   443
  that piece of proof to be treated as @{text invisible} instead of
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   444
  @{text "proof"} (the default), which may be shown or hidden
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   445
  depending on the document setup.  In contrast, ``@{command
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   446
  "by"}~@{text "%visible auto"}'' forces this text to be shown
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   447
  invariably.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   448
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   449
  Explicit tag specifications within a proof apply to all subsequent
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   450
  commands of the same level of nesting.  For example, ``@{command
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   451
  "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   452
  sub-proof to be typeset as @{text visible} (unless some of its parts
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   453
  are tagged differently).
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   454
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   455
  \medskip Command tags merely produce certain markup environments for
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   456
  type-setting.  The meaning of these is determined by {\LaTeX}
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40255
diff changeset
   457
  macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   458
  by the document author.  The Isabelle document preparation tools
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   459
  also provide some high-level options to specify the meaning of
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   460
  arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   461
  parts of the text.  Logic sessions may also specify ``document
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   462
  versions'', where given tags are interpreted in some particular way.
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   463
  Again see \cite{isabelle-sys} for further details.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   464
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   465
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   466
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   467
section {* Draft presentation *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   468
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   469
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   470
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   471
    @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   472
    @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   473
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   474
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   475
  \begin{rail}
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 39689
diff changeset
   476
    ('display_drafts' | 'print_drafts') (name +)
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   477
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   478
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   479
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   480
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   481
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   482
  \item @{command "display_drafts"}~@{text paths} and @{command
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   483
  "print_drafts"}~@{text paths} perform simple output of a given list
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   484
  of raw source files.  Only those symbols that do not require
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   485
  additional {\LaTeX} packages are displayed properly, everything else
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   486
  is left verbatim.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   487
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   488
  \end{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   489
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   490
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   491
end