doc-src/IsarRef/Thy/Document_Preparation.thy
author blanchet
Mon, 19 Apr 2010 18:14:45 +0200
changeset 36230 43d10a494c91
parent 32898 e871d897969c
child 39305 d4fa19eb0822
permissions -rw-r--r--
added warning about inconsistent context to Metis; it makes more sense here than in Sledgehammer, because Sledgehammer is unsound and there's no point in having people panicking about the consistency of their context when their context is in fact consistent
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
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    77
  @{"file" "~~/lib/texinputs/isabelle.sty"} can be redefined according
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
    ;
27049
5072d6c77baa moved header command here;
wenzelm
parents: 27043
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} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   149
    @{antiquotation_def "text"} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   150
    @{antiquotation_def goals} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   151
    @{antiquotation_def subgoals} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   152
    @{antiquotation_def prf} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   153
    @{antiquotation_def full_prf} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   154
    @{antiquotation_def ML} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   155
    @{antiquotation_def ML_type} & : & @{text antiquotation} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   156
    @{antiquotation_def ML_struct} & : & @{text antiquotation} \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   157
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   158
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   159
  The overall content of an Isabelle/Isar theory may alternate between
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   160
  formal and informal text.  The main body consists of formal
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   161
  specification and proof commands, interspersed with markup commands
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   162
  (\secref{sec:markup}) or document comments (\secref{sec:comments}).
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   163
  The argument of markup commands quotes informal text to be printed
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   164
  in the resulting document, but may again refer to formal entities
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   165
  via \emph{document antiquotations}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   166
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   167
  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
   168
  within a text block makes
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   169
  \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
   170
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   171
  Antiquotations usually spare the author tedious typing of logical
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   172
  entities in full detail.  Even more importantly, some degree of
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   173
  consistency-checking between the main body of formal text and its
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   174
  informal explanation is achieved, since terms and types appearing in
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   175
  antiquotations are checked within the current theory or proof
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   176
  context.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   177
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   178
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   179
    atsign lbrace antiquotation rbrace
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   180
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   181
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   182
    antiquotation:
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   183
      'theory' options name |
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   184
      'thm' options styles thmrefs |
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27053
diff changeset
   185
      'lemma' options prop 'by' method |
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   186
      'prop' options styles prop |
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   187
      'term' options styles term |
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   188
      'term_type' options styles term |
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   189
      'typeof' options styles term |
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   190
      'const' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   191
      'abbrev' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   192
      'typ' options type |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   193
      'text' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   194
      'goals' options |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   195
      'subgoals' options |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   196
      'prf' options thmrefs |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   197
      'full\_prf' options thmrefs |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   198
      'ML' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   199
      'ML\_type' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   200
      'ML\_struct' options name
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   201
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   202
    options: '[' (option * ',') ']'
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   203
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   204
    option: name | name '=' name
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   205
    ;
32892
9742f6130f10 corrected syntax diagrams for styles
haftmann
parents: 32891
diff changeset
   206
    styles: '(' (style + ',') ')'
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   207
    ;
32892
9742f6130f10 corrected syntax diagrams for styles
haftmann
parents: 32891
diff changeset
   208
    style: (name +)
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   209
    ;
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   210
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   211
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   212
  Note that the syntax of antiquotations may \emph{not} include source
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   213
  comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   214
  text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   215
  "*"}@{verbatim "}"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   216
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   217
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   218
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   219
  \item @{text "@{theory A}"} prints the name @{text "A"}, which is
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   220
  guaranteed to refer to a valid ancestor theory in the current
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   221
  context.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   222
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   223
  \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
   224
  Full fact expressions are allowed here, including attributes
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   225
  (\secref{sec:syn-att}).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   226
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   227
  \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   228
  "\<phi>"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   229
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   230
  \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   231
  @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27053
diff changeset
   232
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   233
  \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   234
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   235
  \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
   236
  annotated with its type.
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   237
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   238
  \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
   239
  @{text "t"}.
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   240
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   241
  \item @{text "@{const c}"} prints a logical or syntactic constant
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   242
  @{text "c"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   243
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   244
  \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
   245
  @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   246
  \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   247
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   248
  \item @{text "@{text s}"} prints uninterpreted source text @{text
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   249
  s}.  This is particularly useful to print portions of text according
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   250
  to the Isabelle document style, without demanding well-formedness,
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   251
  e.g.\ small pieces of terms that should not be parsed or
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   252
  type-checked yet.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   253
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   254
  \item @{text "@{goals}"} prints the current \emph{dynamic} goal
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   255
  state.  This is mainly for support of tactic-emulation scripts
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   256
  within Isar.  Presentation of goal states does not conform to the
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   257
  idea of human-readable proof documents!
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   258
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   259
  When explaining proofs in detail it is usually better to spell out
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   260
  the reasoning via proper Isar proof commands, instead of peeking at
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   261
  the internal machine configuration.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   262
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   263
  \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   264
  does not print the main goal.
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 "@{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
   267
  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
   268
  requires proof terms to be switched on for the current logic
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   269
  session.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   270
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   271
  \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
   272
  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
   273
  information omitted in the compact proof term, which is denoted by
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   274
  ``@{text _}'' placeholders there.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   275
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   276
  \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   277
  "@{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
   278
  structure, respectively.  The source is printed verbatim.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   279
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   280
  \end{description}
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   281
*}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   282
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   283
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   284
subsubsection {* Styled antiquotations *}
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   285
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   286
text {* The antiquotations @{text thm}, @{text prop} and @{text
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   287
  term} admit an extra \emph{style} specification to modify the
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   288
  printed result.  A style is specified by a name with a possibly
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   289
  empty number of arguments;  multiple styles can be sequenced with
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   290
  commas.  The following standard styles are available:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   291
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   292
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   293
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   294
  \item @{text lhs} extracts the first argument of any application
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   295
  form with at least two arguments --- typically meta-level or
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   296
  object-level equality, or any other binary relation.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   297
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   298
  \item @{text rhs} is like @{text lhs}, but extracts the second
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   299
  argument.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   300
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   301
  \item @{text "concl"} extracts the conclusion @{text C} from a rule
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   302
  in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   303
  
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   304
  \item @{text "prem"} @{text n} extract premise number
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   305
  @{text "n"} from from a rule in Horn-clause
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   306
  normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   307
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   308
  \end{description}
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   309
*}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   310
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   311
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   312
subsubsection {* General options *}
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   313
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   314
text {* The following options are available to tune the printed output
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   315
  of antiquotations.  Note that many of these coincide with global ML
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   316
  flags of the same names.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   317
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   318
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   319
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   320
  \item @{antiquotation_option_def show_types}~@{text "= bool"} and
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   321
  @{antiquotation_option_def show_sorts}~@{text "= bool"} control
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   322
  printing of explicit type and sort constraints.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   323
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   324
  \item @{antiquotation_option_def show_structs}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   325
  controls printing of implicit structures.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   326
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   327
  \item @{antiquotation_option_def long_names}~@{text "= bool"} forces
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   328
  names of types and constants etc.\ to be printed in their fully
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   329
  qualified internal form.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   330
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   331
  \item @{antiquotation_option_def short_names}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   332
  forces names of types and constants etc.\ to be printed unqualified.
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   333
  Note that internalizing the output again in the current context may
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   334
  well yield a different result.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   335
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   336
  \item @{antiquotation_option_def unique_names}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   337
  determines whether the printed version of qualified names should be
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   338
  made sufficiently long to avoid overlap with names declared further
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   339
  back.  Set to @{text false} for more concise output.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   340
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   341
  \item @{antiquotation_option_def eta_contract}~@{text "= bool"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   342
  prints terms in @{text \<eta>}-contracted form.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   343
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   344
  \item @{antiquotation_option_def display}~@{text "= bool"} indicates
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   345
  if the text is to be output as multi-line ``display material'',
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   346
  rather than a small piece of text without line breaks (which is the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   347
  default).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   348
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   349
  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
   350
  the main theory text.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   351
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   352
  \item @{antiquotation_option_def break}~@{text "= bool"} controls
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   353
  line breaks in non-display material.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   354
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   355
  \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   356
  if the output should be enclosed in double quotes.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   357
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   358
  \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   359
  name} to the print mode to be used for presentation.  Note that the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   360
  standard setup for {\LaTeX} output is already present by default,
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   361
  including the modes @{text latex} and @{text xsymbols}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   362
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   363
  \item @{antiquotation_option_def margin}~@{text "= nat"} and
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   364
  @{antiquotation_option_def indent}~@{text "= nat"} change the margin
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   365
  or indentation for pretty printing of display material.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   366
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   367
  \item @{antiquotation_option_def goals_limit}~@{text "= nat"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   368
  determines the maximum number of goals to be printed (for goal-based
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   369
  antiquotation).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   370
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   371
  \item @{antiquotation_option_def source}~@{text "= bool"} prints the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   372
  original source text of the antiquotation arguments, rather than its
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   373
  internal representation.  Note that formal checking of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   374
  @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   375
  enabled; use the @{antiquotation "text"} antiquotation for unchecked
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   376
  output.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   377
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   378
  Regular @{text "term"} and @{text "typ"} antiquotations with @{text
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   379
  "source = false"} involve a full round-trip from the original source
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   380
  to an internalized logical entity back to a source form, according
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   381
  to the syntax of the current context.  Thus the printed output is
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   382
  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
   383
  as the underlying theory is changed later on.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   384
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   385
  In contrast, @{antiquotation_option_def source}~@{text "= true"}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   386
  admits direct printing of the given source text, with the desirable
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   387
  well-formedness check in the background, but without modification of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   388
  the printed text.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   389
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   390
  \end{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   391
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   392
  For boolean flags, ``@{text "name = true"}'' may be abbreviated as
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   393
  ``@{text name}''.  All of the above flags are disabled by default,
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   394
  unless changed from ML, say in the @{verbatim "ROOT.ML"} of the
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   395
  logic session.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   396
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   397
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   398
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   399
section {* Markup via command tags \label{sec:tags} *}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   400
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   401
text {* Each Isabelle/Isar command may be decorated by additional
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   402
  presentation tags, to indicate some modification in the way it is
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   403
  printed in the document.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   404
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   405
  \indexouternonterm{tags}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   406
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   407
    tags: ( tag * )
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   408
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   409
    tag: '\%' (ident | string)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   410
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   411
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   412
  Some tags are pre-declared for certain classes of commands, serving
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   413
  as default markup if no tags are given in the text:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   414
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   415
  \medskip
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   416
  \begin{tabular}{ll}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   417
    @{text "theory"} & theory begin/end \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   418
    @{text "proof"} & all proof commands \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   419
    @{text "ML"} & all commands involving ML code \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   420
  \end{tabular}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   421
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   422
  \medskip The Isabelle document preparation system
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   423
  \cite{isabelle-sys} allows tagged command regions to be presented
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   424
  specifically, e.g.\ to fold proof texts, or drop parts of the text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   425
  completely.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   426
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   427
  For example ``@{command "by"}~@{text "%invisible auto"}'' causes
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   428
  that piece of proof to be treated as @{text invisible} instead of
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   429
  @{text "proof"} (the default), which may be shown or hidden
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   430
  depending on the document setup.  In contrast, ``@{command
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   431
  "by"}~@{text "%visible auto"}'' forces this text to be shown
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   432
  invariably.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   433
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   434
  Explicit tag specifications within a proof apply to all subsequent
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   435
  commands of the same level of nesting.  For example, ``@{command
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   436
  "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   437
  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
   438
  are tagged differently).
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   439
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   440
  \medskip Command tags merely produce certain markup environments for
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   441
  type-setting.  The meaning of these is determined by {\LaTeX}
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   442
  macros, as defined in @{"file" "~~/lib/texinputs/isabelle.sty"} or
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   443
  by the document author.  The Isabelle document preparation tools
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   444
  also provide some high-level options to specify the meaning of
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   445
  arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   446
  parts of the text.  Logic sessions may also specify ``document
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   447
  versions'', where given tags are interpreted in some particular way.
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   448
  Again see \cite{isabelle-sys} for further details.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   449
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   450
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   451
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   452
section {* Draft presentation *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   453
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   454
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   455
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   456
    @{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
   457
    @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   458
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   459
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   460
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   461
    ('display\_drafts' | 'print\_drafts') (name +)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   462
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   463
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   464
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   465
  \begin{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   466
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   467
  \item @{command "display_drafts"}~@{text paths} and @{command
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   468
  "print_drafts"}~@{text paths} perform simple output of a given list
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   469
  of raw source files.  Only those symbols that do not require
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   470
  additional {\LaTeX} packages are displayed properly, everything else
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   471
  is left verbatim.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   472
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   473
  \end{description}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   474
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   475
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   476
end