src/Doc/Isar_Ref/Document_Preparation.thy
author wenzelm
Wed, 21 Oct 2015 00:23:11 +0200
changeset 61494 63b18f758874
parent 61493 0debd22f0c0e
child 61503 28e788ca2c5d
permissions -rw-r--r--
proper spaces around @{text};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     1
theory Document_Preparation
42651
e3fdb7c96be5 formal Base theory;
wenzelm
parents: 42626
diff changeset
     2
imports Base Main
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     3
begin
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     4
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
     5
chapter \<open>Document preparation \label{ch:document-prep}\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
     7
text \<open>Isabelle/Isar provides a simple document preparation system
54346
wenzelm
parents: 53982
diff changeset
     8
  based on {PDF-\LaTeX}, with support for hyperlinks and bookmarks
wenzelm
parents: 53982
diff changeset
     9
  within that format.  This allows to produce papers, books, theses
wenzelm
parents: 53982
diff changeset
    10
  etc.\ from Isabelle theory sources.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    11
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61474
diff changeset
    12
  {\LaTeX} output is generated while processing a \<^emph>\<open>session\<close> in
e467ae7aa808 more control symbols;
wenzelm
parents: 61474
diff changeset
    13
  batch mode, as explained in the \<^emph>\<open>The Isabelle System Manual\<close>
60270
wenzelm
parents: 59937
diff changeset
    14
  @{cite "isabelle-system"}.  The main Isabelle tools to get started with
54346
wenzelm
parents: 53982
diff changeset
    15
  document preparation are @{tool_ref mkroot} and @{tool_ref build}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    16
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58069
diff changeset
    17
  The classic Isabelle/HOL tutorial @{cite "isabelle-hol-book"} also
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
    18
  explains some aspects of theory presentation.\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    19
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    20
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
    21
section \<open>Markup commands \label{sec:markup}\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    22
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
    23
text \<open>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    24
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    25
    @{command_def "chapter"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    26
    @{command_def "section"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    27
    @{command_def "subsection"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    28
    @{command_def "subsubsection"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    29
    @{command_def "paragraph"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    30
    @{command_def "subparagraph"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    31
    @{command_def "text"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    32
    @{command_def "txt"} & : & \<open>any \<rightarrow> any\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    33
    @{command_def "text_raw"} & : & \<open>any \<rightarrow> any\<close> \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    34
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    35
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    36
  Markup commands provide a structured way to insert text into the
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    37
  document generated from a theory.  Each markup command takes a
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    38
  single @{syntax text} argument, which is passed as argument to a
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    39
  corresponding {\LaTeX} macro.  The default macros provided by
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40255
diff changeset
    40
  @{file "~~/lib/texinputs/isabelle.sty"} can be redefined according
28747
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    41
  to the needs of the underlying document and {\LaTeX} styles.
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    42
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    43
  Note that formal comments (\secref{sec:comments}) are similar to
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    44
  markup commands, but have a different status within Isabelle/Isar
ec3424dd17bc tuned "Markup commands";
wenzelm
parents: 28746
diff changeset
    45
  syntax.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    46
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
    47
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
    48
    (@@{command chapter} | @@{command section} | @@{command subsection} |
61463
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    49
      @@{command subsubsection} | @@{command paragraph} | @@{command subparagraph} |
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    50
      @@{command text} | @@{command txt} | @@{command text_raw}) @{syntax text}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
    51
  \<close>}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    52
61463
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    53
  \<^descr> @{command chapter}, @{command section}, @{command subsection} etc.\ mark
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    54
  section headings within the theory source. This works in any context, even
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    55
  before the initial @{command theory} command. The corresponding {\LaTeX}
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    56
  macros are @{verbatim \<open>\isamarkupchapter\<close>}, @{verbatim
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    57
  \<open>\isamarkupsection\<close>}, @{verbatim \<open>\isamarkupsubsection\<close>} etc.\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    58
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
    59
  \<^descr> @{command text} and @{command txt} specify paragraphs of plain text.
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58868
diff changeset
    60
  This corresponds to a {\LaTeX} environment @{verbatim
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    61
  \<open>\begin{isamarkuptext}\<close>} \<open>\<dots>\<close> @{verbatim \<open>\end{isamarkuptext}\<close>}
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58868
diff changeset
    62
  etc.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    63
61463
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    64
  \<^descr> @{command text_raw} is similar to @{command text}, but without
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    65
  any surrounding markup environment. This allows to inject arbitrary
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    66
  {\LaTeX} source into the generated document.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    67
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    68
61463
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    69
  All text passed to any of the above markup commands may refer to formal
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61474
diff changeset
    70
  entities via \<^emph>\<open>document antiquotations\<close>, see also \secref{sec:antiq}.
61463
8e46cea6a45a added 'paragraph', 'subparagraph';
wenzelm
parents: 61458
diff changeset
    71
  These are interpreted in the present theory or proof context.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    72
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
    73
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
    74
  The proof markup commands closely resemble those for theory
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    75
  specifications, but have a different formal status and produce
58868
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58725
diff changeset
    76
  different {\LaTeX} macros.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
    77
\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    78
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    79
60286
wenzelm
parents: 60270
diff changeset
    80
section \<open>Document antiquotations \label{sec:antiq}\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    81
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
    82
text \<open>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    83
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    84
    @{antiquotation_def "theory"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    85
    @{antiquotation_def "thm"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    86
    @{antiquotation_def "lemma"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    87
    @{antiquotation_def "prop"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    88
    @{antiquotation_def "term"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    89
    @{antiquotation_def term_type} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    90
    @{antiquotation_def typeof} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    91
    @{antiquotation_def const} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    92
    @{antiquotation_def abbrev} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    93
    @{antiquotation_def typ} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    94
    @{antiquotation_def type} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    95
    @{antiquotation_def class} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    96
    @{antiquotation_def "text"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    97
    @{antiquotation_def goals} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    98
    @{antiquotation_def subgoals} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
    99
    @{antiquotation_def prf} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   100
    @{antiquotation_def full_prf} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   101
    @{antiquotation_def ML} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   102
    @{antiquotation_def ML_op} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   103
    @{antiquotation_def ML_type} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   104
    @{antiquotation_def ML_structure} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   105
    @{antiquotation_def ML_functor} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   106
    @{antiquotation_def emph} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   107
    @{antiquotation_def bold} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   108
    @{antiquotation_def verbatim} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   109
    @{antiquotation_def "file"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   110
    @{antiquotation_def "url"} & : & \<open>antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   111
    @{antiquotation_def "cite"} & : & \<open>antiquotation\<close> \\
61494
63b18f758874 proper spaces around @{text};
wenzelm
parents: 61493
diff changeset
   112
    @{command_def "print_antiquotations"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   113
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   114
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   115
  The overall content of an Isabelle/Isar theory may alternate between
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   116
  formal and informal text.  The main body consists of formal
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   117
  specification and proof commands, interspersed with markup commands
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   118
  (\secref{sec:markup}) or document comments (\secref{sec:comments}).
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   119
  The argument of markup commands quotes informal text to be printed
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   120
  in the resulting document, but may again refer to formal entities
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61474
diff changeset
   121
  via \<^emph>\<open>document antiquotations\<close>.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   122
58724
e5f809f52f26 more antiquotations;
wenzelm
parents: 58716
diff changeset
   123
  For example, embedding @{verbatim \<open>@{term [show_types] "f x = a + x"}\<close>}
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   124
  within a text block makes
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   125
  \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
   126
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   127
  Antiquotations usually spare the author tedious typing of logical
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   128
  entities in full detail.  Even more importantly, some degree of
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   129
  consistency-checking between the main body of formal text and its
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   130
  informal explanation is achieved, since terms and types appearing in
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   131
  antiquotations are checked within the current theory or proof
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   132
  context.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   133
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   134
  \<^medskip> Antiquotations are in general written as @{verbatim "@{"}\<open>name\<close>~@{verbatim "["}\<open>options\<close>@{verbatim "]"}~\<open>arguments\<close>@{verbatim "}"}. The short form @{verbatim \<open>\\<close>}@{verbatim
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   135
  "<^"}\<open>name\<close>@{verbatim ">"}\<open>\<open>argument_content\<close>\<close> (without
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   136
  surrounding @{verbatim "@{"}\<open>\<dots>\<close>@{verbatim "}"}) works for a single
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   137
  argument that is a cartouche.
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   138
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   139
  Omitting the control symbol is also possible: a cartouche without special
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   140
  decoration is equivalent to @{verbatim \<open>\\<close>}@{verbatim
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   141
  "<^cartouche>"}\<open>\<open>argument_content\<close>\<close>, which is equivalent to
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   142
  @{verbatim "@{cartouche"}~\<open>\<open>argument_content\<close>\<close>@{verbatim "}"}. The
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   143
  special name @{antiquotation_def cartouche} is defined in the context:
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   144
  Isabelle/Pure introduces that as an alias to @{antiquotation_ref text}
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   145
  (see below). Consequently, \<open>\<open>foo_bar + baz \<le> bazar\<close>\<close> prints literal
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   146
  quasi-formal text (unchecked).
61472
6458760261ca more documentation;
wenzelm
parents: 61463
diff changeset
   147
61474
6cc07122ca14 clarified;
wenzelm
parents: 61473
diff changeset
   148
  \begingroup
6cc07122ca14 clarified;
wenzelm
parents: 61473
diff changeset
   149
  \def\isasymcontrolstart{\isatt{\isacharbackslash\isacharless\isacharcircum}}
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59783
diff changeset
   150
  @{rail \<open>
61472
6458760261ca more documentation;
wenzelm
parents: 61463
diff changeset
   151
    @{syntax_def antiquotation}:
61474
6cc07122ca14 clarified;
wenzelm
parents: 61473
diff changeset
   152
      '@{' antiquotation_body '}' |
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   153
      '\<controlstart>' @{syntax_ref name} '>' @{syntax_ref cartouche} |
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   154
      @{syntax_ref cartouche}
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   155
    ;
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   156
    options: '[' (option * ',') ']'
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   157
    ;
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   158
    option: @{syntax name} | @{syntax name} '=' @{syntax name}
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   159
    ;
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59783
diff changeset
   160
  \<close>}
61474
6cc07122ca14 clarified;
wenzelm
parents: 61473
diff changeset
   161
  \endgroup
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59783
diff changeset
   162
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   163
  Note that the syntax of antiquotations may \<^emph>\<open>not\<close> include source
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   164
  comments @{verbatim "(*"}~\<open>\<dots>\<close>~@{verbatim "*)"} nor verbatim
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   165
  text @{verbatim "{*"}~\<open>\<dots>\<close>~@{verbatim "*}"}.
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   166
43618
wenzelm
parents: 43613
diff changeset
   167
  %% FIXME less monolithic presentation, move to individual sections!?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   168
  @{rail \<open>
61472
6458760261ca more documentation;
wenzelm
parents: 61463
diff changeset
   169
    @{syntax_def antiquotation_body}:
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   170
      (@@{antiquotation text} | @@{antiquotation cartouche}) options @{syntax text} |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   171
      @@{antiquotation theory} options @{syntax name} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   172
      @@{antiquotation thm} options styles @{syntax thmrefs} |
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   173
      @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   174
      @@{antiquotation prop} options styles @{syntax prop} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   175
      @@{antiquotation term} options styles @{syntax term} |
43618
wenzelm
parents: 43613
diff changeset
   176
      @@{antiquotation (HOL) value} options styles @{syntax term} |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   177
      @@{antiquotation term_type} options styles @{syntax term} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   178
      @@{antiquotation typeof} options styles @{syntax term} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   179
      @@{antiquotation const} options @{syntax term} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   180
      @@{antiquotation abbrev} options @{syntax term} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   181
      @@{antiquotation typ} options @{syntax type} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   182
      @@{antiquotation type} options @{syntax name} |
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   183
      @@{antiquotation class} options @{syntax name}
46261
b03897da3c90 document antiquotations for ML infix operators;
wenzelm
parents: 43618
diff changeset
   184
    ;
b03897da3c90 document antiquotations for ML infix operators;
wenzelm
parents: 43618
diff changeset
   185
    @{syntax antiquotation}:
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   186
      @@{antiquotation goals} options |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   187
      @@{antiquotation subgoals} options |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   188
      @@{antiquotation prf} options @{syntax thmrefs} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   189
      @@{antiquotation full_prf} options @{syntax thmrefs} |
58069
0255436b3d85 more liberal embedded "text", which includes cartouches;
wenzelm
parents: 56594
diff changeset
   190
      @@{antiquotation ML} options @{syntax text} |
0255436b3d85 more liberal embedded "text", which includes cartouches;
wenzelm
parents: 56594
diff changeset
   191
      @@{antiquotation ML_op} options @{syntax text} |
0255436b3d85 more liberal embedded "text", which includes cartouches;
wenzelm
parents: 56594
diff changeset
   192
      @@{antiquotation ML_type} options @{syntax text} |
0255436b3d85 more liberal embedded "text", which includes cartouches;
wenzelm
parents: 56594
diff changeset
   193
      @@{antiquotation ML_structure} options @{syntax text} |
0255436b3d85 more liberal embedded "text", which includes cartouches;
wenzelm
parents: 56594
diff changeset
   194
      @@{antiquotation ML_functor} options @{syntax text} |
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   195
      @@{antiquotation emph} options @{syntax text} |
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   196
      @@{antiquotation bold} options @{syntax text} |
58716
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58618
diff changeset
   197
      @@{antiquotation verbatim} options @{syntax text} |
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   198
      @@{antiquotation "file"} options @{syntax name} |
54705
0dff3326d12a provide @{file_unchecked} in Isabelle/Pure;
wenzelm
parents: 54702
diff changeset
   199
      @@{antiquotation file_unchecked} options @{syntax name} |
58593
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   200
      @@{antiquotation url} options @{syntax name} |
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   201
      @@{antiquotation cite} options @{syntax cartouche}? (@{syntax name} + @'and')
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   202
    ;
32892
9742f6130f10 corrected syntax diagrams for styles
haftmann
parents: 32891
diff changeset
   203
    styles: '(' (style + ',') ')'
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   204
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   205
    style: (@{syntax name} +)
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   206
    ;
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   207
    @@{command print_antiquotations} ('!'?)
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   208
  \<close>}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   209
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   210
  \<^descr> \<open>@{text s}\<close> prints uninterpreted source text \<open>s\<close>.  This is particularly useful to print portions of text according
61491
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   211
  to the Isabelle document style, without demanding well-formedness,
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   212
  e.g.\ small pieces of terms that should not be parsed or
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   213
  type-checked yet.
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   214
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   215
  It is also possible to write this in the short form \<open>\<open>s\<close>\<close> without any
97261e6c1d42 another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents: 61479
diff changeset
   216
  further decoration.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   217
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   218
  \<^descr> \<open>@{theory A}\<close> prints the name \<open>A\<close>, which is
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   219
  guaranteed to refer to a valid ancestor theory in the current
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   220
  context.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   221
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   222
  \<^descr> \<open>@{thm a\<^sub>1 \<dots> a\<^sub>n}\<close> prints theorems \<open>a\<^sub>1 \<dots> a\<^sub>n\<close>.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   223
  Full fact expressions are allowed here, including attributes
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   224
  (\secref{sec:syn-att}).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   225
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   226
  \<^descr> \<open>@{prop \<phi>}\<close> prints a well-typed proposition \<open>\<phi>\<close>.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   227
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   228
  \<^descr> \<open>@{lemma \<phi> by m}\<close> proves a well-typed proposition
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   229
  \<open>\<phi>\<close> by method \<open>m\<close> and prints the original \<open>\<phi>\<close>.
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27053
diff changeset
   230
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   231
  \<^descr> \<open>@{term t}\<close> prints a well-typed term \<open>t\<close>.
43613
7afbaf5a406b adding a minimalistic documentation of the value antiquotation in the Isar reference manual
bulwahn
parents: 42936
diff changeset
   232
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   233
  \<^descr> \<open>@{value t}\<close> evaluates a term \<open>t\<close> and prints
43618
wenzelm
parents: 43613
diff changeset
   234
  its result, see also @{command_ref (HOL) value}.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   235
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   236
  \<^descr> \<open>@{term_type t}\<close> prints a well-typed term \<open>t\<close>
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   237
  annotated with its type.
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   238
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   239
  \<^descr> \<open>@{typeof t}\<close> prints the type of a well-typed term
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   240
  \<open>t\<close>.
32898
e871d897969c term styles also cover antiquotations term_type and typeof
haftmann
parents: 32892
diff changeset
   241
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   242
  \<^descr> \<open>@{const c}\<close> prints a logical or syntactic constant
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   243
  \<open>c\<close>.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   244
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   245
  \<^descr> \<open>@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}\<close> prints a constant abbreviation
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   246
  \<open>c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs\<close> as defined in the current context.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   247
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   248
  \<^descr> \<open>@{typ \<tau>}\<close> prints a well-formed type \<open>\<tau>\<close>.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   249
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   250
  \<^descr> \<open>@{type \<kappa>}\<close> prints a (logical or syntactic) type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   251
    constructor \<open>\<kappa>\<close>.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   252
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   253
  \<^descr> \<open>@{class c}\<close> prints a class \<open>c\<close>.
39305
d4fa19eb0822 'class' and 'type' are now antiquoations by default
haftmann
parents: 32898
diff changeset
   254
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   255
  \<^descr> \<open>@{goals}\<close> prints the current \<^emph>\<open>dynamic\<close> goal
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   256
  state.  This is mainly for support of tactic-emulation scripts
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   257
  within Isar.  Presentation of goal states does not conform to the
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   258
  idea of human-readable proof documents!
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   259
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   260
  When explaining proofs in detail it is usually better to spell out
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   261
  the reasoning via proper Isar proof commands, instead of peeking at
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   262
  the internal machine configuration.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   263
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   264
  \<^descr> \<open>@{subgoals}\<close> is similar to \<open>@{goals}\<close>, but
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   265
  does not print the main goal.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   266
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   267
  \<^descr> \<open>@{prf a\<^sub>1 \<dots> a\<^sub>n}\<close> prints the (compact) proof terms
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   268
  corresponding to the theorems \<open>a\<^sub>1 \<dots> a\<^sub>n\<close>. Note that this
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   269
  requires proof terms to be switched on for the current logic
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   270
  session.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   271
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   272
  \<^descr> \<open>@{full_prf a\<^sub>1 \<dots> a\<^sub>n}\<close> is like \<open>@{prf a\<^sub>1 \<dots>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   273
  a\<^sub>n}\<close>, but prints the full proof terms, i.e.\ also displays
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28750
diff changeset
   274
  information omitted in the compact proof term, which is denoted by
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   275
  ``\<open>_\<close>'' placeholders there.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   276
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   277
  \<^descr> \<open>@{ML s}\<close>, \<open>@{ML_op s}\<close>, \<open>@{ML_type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   278
  s}\<close>, \<open>@{ML_structure s}\<close>, and \<open>@{ML_functor s}\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   279
  check text \<open>s\<close> as ML value, infix operator, type, structure,
55837
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55120
diff changeset
   280
  and functor respectively.  The source is printed verbatim.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   281
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   282
  \<^descr> \<open>@{emph s}\<close> prints document source recursively, with {\LaTeX}
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   283
  markup @{verbatim \<open>\emph{\<close>}\<open>\<dots>\<close>@{verbatim \<open>}\<close>}.
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   284
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   285
  \<^descr> \<open>@{bold s}\<close> prints document source recursively, with {\LaTeX}
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   286
  markup @{verbatim \<open>\textbf{\<close>}\<open>\<dots>\<close>@{verbatim \<open>}\<close>}.
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   287
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   288
  \<^descr> \<open>@{verbatim s}\<close> prints uninterpreted source text literally
58716
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58618
diff changeset
   289
  as ASCII characters, using some type-writer font style.
23a380cc45f4 official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents: 58618
diff changeset
   290
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   291
  \<^descr> \<open>@{file path}\<close> checks that \<open>path\<close> refers to a
40801
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   292
  file (or directory) and prints it verbatim.
6cfacec435e6 added document antiquotation @{file};
wenzelm
parents: 40800
diff changeset
   293
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   294
  \<^descr> \<open>@{file_unchecked path}\<close> is like \<open>@{file
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   295
  path}\<close>, but does not check the existence of the \<open>path\<close>
54705
0dff3326d12a provide @{file_unchecked} in Isabelle/Pure;
wenzelm
parents: 54702
diff changeset
   296
  within the file-system.
0dff3326d12a provide @{file_unchecked} in Isabelle/Pure;
wenzelm
parents: 54702
diff changeset
   297
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   298
  \<^descr> \<open>@{url name}\<close> produces markup for the given URL, which
54702
3daeba5130f0 added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
wenzelm
parents: 54346
diff changeset
   299
  results in an active hyperlink within the text.
3daeba5130f0 added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
wenzelm
parents: 54346
diff changeset
   300
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   301
  \<^descr> \<open>@{cite name}\<close> produces a citation @{verbatim
58724
e5f809f52f26 more antiquotations;
wenzelm
parents: 58716
diff changeset
   302
  \<open>\cite{name}\<close>} in {\LaTeX}, where the name refers to some Bib{\TeX}
58593
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   303
  database entry.
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   304
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   305
  The variant \<open>@{cite \<open>opt\<close> name}\<close> produces @{verbatim
58724
e5f809f52f26 more antiquotations;
wenzelm
parents: 58716
diff changeset
   306
  \<open>\cite[opt]{name}\<close>} with some free-form optional argument. Multiple names
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   307
  are output with commas, e.g. \<open>@{cite foo \<AND> bar}\<close> becomes
58724
e5f809f52f26 more antiquotations;
wenzelm
parents: 58716
diff changeset
   308
  @{verbatim \<open>\cite{foo,bar}\<close>}.
58593
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   309
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   310
  The {\LaTeX} macro name is determined by the antiquotation option
51adee3ace7b documentation of @{cite} and cite_macro;
wenzelm
parents: 58552
diff changeset
   311
  @{antiquotation_option_def cite_macro}, or the configuration option
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   312
  @{attribute cite_macro} in the context. For example, \<open>@{cite
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   313
  [cite_macro = nocite] foobar}\<close> produces @{verbatim \<open>\nocite{foobar}\<close>}.
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   314
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   315
  \<^descr> @{command "print_antiquotations"} prints all document antiquotations
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   316
  that are defined in the current context; the ``\<open>!\<close>'' option
61473
34d1913f0b20 clarified control antiquotations: decode control symbol to get name;
wenzelm
parents: 61472
diff changeset
   317
  indicates extra verbosity.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   318
\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   319
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   320
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   321
subsection \<open>Styled antiquotations\<close>
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   322
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   323
text \<open>The antiquotations \<open>thm\<close>, \<open>prop\<close> and \<open>term\<close> admit an extra \<^emph>\<open>style\<close> specification to modify the
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   324
  printed result.  A style is specified by a name with a possibly
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   325
  empty number of arguments;  multiple styles can be sequenced with
d403b99287ff new generalized concept for term styles
haftmann
parents: 30397
diff changeset
   326
  commas.  The following standard styles are available:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   327
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   328
  \<^descr> \<open>lhs\<close> extracts the first argument of any application
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   329
  form with at least two arguments --- typically meta-level or
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   330
  object-level equality, or any other binary relation.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   331
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   332
  \<^descr> \<open>rhs\<close> is like \<open>lhs\<close>, but extracts the second
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   333
  argument.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   334
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   335
  \<^descr> \<open>concl\<close> extracts the conclusion \<open>C\<close> from a rule
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   336
  in Horn-clause normal form \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C\<close>.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   337
  
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   338
  \<^descr> \<open>prem\<close> \<open>n\<close> extract premise number
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   339
  \<open>n\<close> from from a rule in Horn-clause
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   340
  normal form \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C\<close>
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   341
\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   342
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   343
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   344
subsection \<open>General options\<close>
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   345
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   346
text \<open>The following options are available to tune the printed output
54346
wenzelm
parents: 53982
diff changeset
   347
  of antiquotations.  Note that many of these coincide with system and
wenzelm
parents: 53982
diff changeset
   348
  configuration options of the same names.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   349
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   350
  \<^descr> @{antiquotation_option_def show_types}~\<open>= bool\<close> and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   351
  @{antiquotation_option_def show_sorts}~\<open>= bool\<close> control
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   352
  printing of explicit type and sort constraints.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   353
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   354
  \<^descr> @{antiquotation_option_def show_structs}~\<open>= bool\<close>
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   355
  controls printing of implicit structures.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   356
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   357
  \<^descr> @{antiquotation_option_def show_abbrevs}~\<open>= bool\<close>
40879
ca132ef44944 configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents: 40801
diff changeset
   358
  controls folding of abbreviations.
ca132ef44944 configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents: 40801
diff changeset
   359
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   360
  \<^descr> @{antiquotation_option_def names_long}~\<open>= bool\<close> forces
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   361
  names of types and constants etc.\ to be printed in their fully
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   362
  qualified internal form.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   363
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   364
  \<^descr> @{antiquotation_option_def names_short}~\<open>= bool\<close>
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   365
  forces names of types and constants etc.\ to be printed unqualified.
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   366
  Note that internalizing the output again in the current context may
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   367
  well yield a different result.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   368
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   369
  \<^descr> @{antiquotation_option_def names_unique}~\<open>= bool\<close>
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   370
  determines whether the printed version of qualified names should be
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   371
  made sufficiently long to avoid overlap with names declared further
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   372
  back.  Set to \<open>false\<close> for more concise output.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   373
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   374
  \<^descr> @{antiquotation_option_def eta_contract}~\<open>= bool\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   375
  prints terms in \<open>\<eta>\<close>-contracted form.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   376
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   377
  \<^descr> @{antiquotation_option_def display}~\<open>= bool\<close> indicates
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   378
  if the text is to be output as multi-line ``display material'',
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   379
  rather than a small piece of text without line breaks (which is the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   380
  default).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   381
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   382
  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
   383
  the main theory text.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   384
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   385
  \<^descr> @{antiquotation_option_def break}~\<open>= bool\<close> controls
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   386
  line breaks in non-display material.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   387
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   388
  \<^descr> @{antiquotation_option_def quotes}~\<open>= bool\<close> indicates
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   389
  if the output should be enclosed in double quotes.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   390
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   391
  \<^descr> @{antiquotation_option_def mode}~\<open>= name\<close> adds \<open>name\<close> to the print mode to be used for presentation.  Note that the
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   392
  standard setup for {\LaTeX} output is already present by default,
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   393
  including the modes \<open>latex\<close> and \<open>xsymbols\<close>.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   394
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   395
  \<^descr> @{antiquotation_option_def margin}~\<open>= nat\<close> and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   396
  @{antiquotation_option_def indent}~\<open>= nat\<close> change the margin
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   397
  or indentation for pretty printing of display material.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   398
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   399
  \<^descr> @{antiquotation_option_def goals_limit}~\<open>= nat\<close>
51960
61ac1efe02c3 option "goals_limit", with more uniform description;
wenzelm
parents: 51057
diff changeset
   400
  determines the maximum number of subgoals to be printed (for goal-based
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   401
  antiquotation).
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   402
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   403
  \<^descr> @{antiquotation_option_def source}~\<open>= bool\<close> prints the
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   404
  original source text of the antiquotation arguments, rather than its
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   405
  internal representation.  Note that formal checking of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   406
  @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   407
  enabled; use the @{antiquotation "text"} antiquotation for unchecked
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   408
  output.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   409
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   410
  Regular \<open>term\<close> and \<open>typ\<close> antiquotations with \<open>source = false\<close> involve a full round-trip from the original source
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   411
  to an internalized logical entity back to a source form, according
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   412
  to the syntax of the current context.  Thus the printed output is
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   413
  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
   414
  as the underlying theory is changed later on.
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   415
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   416
  In contrast, @{antiquotation_option source}~\<open>= true\<close>
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   417
  admits direct printing of the given source text, with the desirable
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   418
  well-formedness check in the background, but without modification of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   419
  the printed text.
28749
99f6da3bbbf7 renamed "formal comments" to "document comments";
wenzelm
parents: 28747
diff changeset
   420
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   421
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   422
  For Boolean flags, ``\<open>name = true\<close>'' may be abbreviated as
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   423
  ``\<open>name\<close>''.  All of the above flags are disabled by default,
51057
a22b134f862e updated explanations of document preparation;
wenzelm
parents: 48985
diff changeset
   424
  unless changed specifically for a logic session in the corresponding
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   425
  @{verbatim "ROOT"} file.
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   426
\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   427
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   428
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   429
section \<open>Markup via command tags \label{sec:tags}\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   430
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   431
text \<open>Each Isabelle/Isar command may be decorated by additional
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   432
  presentation tags, to indicate some modification in the way it is
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   433
  printed in the document.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   434
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   435
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   436
    @{syntax_def tags}: ( tag * )
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   437
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40879
diff changeset
   438
    tag: '%' (@{syntax ident} | @{syntax string})
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   439
  \<close>}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   440
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   441
  Some tags are pre-declared for certain classes of commands, serving
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   442
  as default markup if no tags are given in the text:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   443
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   444
  \<^medskip>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   445
  \begin{tabular}{ll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   446
    \<open>theory\<close> & theory begin/end \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   447
    \<open>proof\<close> & all proof commands \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   448
    \<open>ML\<close> & all commands involving ML code \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   449
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   450
  \<^medskip>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   451
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   452
  The Isabelle document preparation system
60270
wenzelm
parents: 59937
diff changeset
   453
  @{cite "isabelle-system"} allows tagged command regions to be presented
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   454
  specifically, e.g.\ to fold proof texts, or drop parts of the text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   455
  completely.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   456
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   457
  For example ``@{command "by"}~\<open>%invisible auto\<close>'' causes
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   458
  that piece of proof to be treated as \<open>invisible\<close> instead of
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   459
  \<open>proof\<close> (the default), which may be shown or hidden
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   460
  depending on the document setup.  In contrast, ``@{command
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   461
  "by"}~\<open>%visible auto\<close>'' forces this text to be shown
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   462
  invariably.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   463
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   464
  Explicit tag specifications within a proof apply to all subsequent
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   465
  commands of the same level of nesting.  For example, ``@{command
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   466
  "proof"}~\<open>%visible \<dots>\<close>~@{command "qed"}'' forces the whole
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   467
  sub-proof to be typeset as \<open>visible\<close> (unless some of its parts
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   468
  are tagged differently).
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   469
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   470
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   471
  Command tags merely produce certain markup environments for
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   472
  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
   473
  macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
28750
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   474
  by the document author.  The Isabelle document preparation tools
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   475
  also provide some high-level options to specify the meaning of
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   476
  arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   477
  parts of the text.  Logic sessions may also specify ``document
1ff7fff6a170 updated section "Markup via command tags";
wenzelm
parents: 28749
diff changeset
   478
  versions'', where given tags are interpreted in some particular way.
60270
wenzelm
parents: 59937
diff changeset
   479
  Again see @{cite "isabelle-system"} for further details.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   480
\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   481
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   482
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   483
section \<open>Railroad diagrams\<close>
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   484
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   485
text \<open>
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   486
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   487
    @{antiquotation_def "rail"} & : & \<open>antiquotation\<close> \\
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   488
  \end{matharray}
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   489
55113
wenzelm
parents: 55112
diff changeset
   490
  @{rail \<open>
59937
6eccb133d4e6 clarified rail syntax;
wenzelm
parents: 59917
diff changeset
   491
    'rail' @{syntax text}
55113
wenzelm
parents: 55112
diff changeset
   492
  \<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   493
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   494
  The @{antiquotation rail} antiquotation allows to include syntax
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   495
  diagrams into Isabelle documents.  {\LaTeX} requires the style file
59116
77351f2051f5 proper railsetup;
wenzelm
parents: 58999
diff changeset
   496
  @{file "~~/lib/texinputs/railsetup.sty"}, which can be used via
77351f2051f5 proper railsetup;
wenzelm
parents: 58999
diff changeset
   497
  @{verbatim \<open>\usepackage{railsetup}\<close>} in @{verbatim "root.tex"}, for
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   498
  example.
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   499
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   500
  The rail specification language is quoted here as Isabelle @{syntax
55120
wenzelm
parents: 55113
diff changeset
   501
  string} or text @{syntax "cartouche"}; it has its own grammar given
wenzelm
parents: 55113
diff changeset
   502
  below.
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   503
55120
wenzelm
parents: 55113
diff changeset
   504
  \begingroup
61474
6cc07122ca14 clarified;
wenzelm
parents: 61473
diff changeset
   505
  \def\isasymnewline{\isatt{\isacharbackslash\isacharless newline\isachargreater}}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   506
  @{rail \<open>
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   507
  rule? + ';'
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   508
  ;
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   509
  rule: ((identifier | @{syntax antiquotation}) ':')? body
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   510
  ;
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   511
  body: concatenation + '|'
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   512
  ;
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   513
  concatenation: ((atom '?'?) +) (('*' | '+') atom?)?
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   514
  ;
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   515
  atom: '(' body? ')' | identifier |
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   516
    '@'? (string | @{syntax antiquotation}) |
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   517
    '\<newline>'
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   518
  \<close>}
55120
wenzelm
parents: 55113
diff changeset
   519
  \endgroup
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   520
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   521
  The lexical syntax of \<open>identifier\<close> coincides with that of
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   522
  @{syntax ident} in regular Isabelle syntax, but \<open>string\<close> uses
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   523
  single quotes instead of double quotes of the standard @{syntax
55113
wenzelm
parents: 55112
diff changeset
   524
  string} category.
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   525
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   526
  Each \<open>rule\<close> defines a formal language (with optional name),
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   527
  using a notation that is similar to EBNF or regular expressions with
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   528
  recursion.  The meaning and visual appearance of these rail language
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   529
  elements is illustrated by the following representative examples.
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   530
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   531
  \<^item> Empty @{verbatim "()"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   532
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   533
  @{rail \<open>()\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   534
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   535
  \<^item> Nonterminal @{verbatim "A"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   536
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   537
  @{rail \<open>A\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   538
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   539
  \<^item> Nonterminal via Isabelle antiquotation
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   540
  @{verbatim "@{syntax method}"}
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   541
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   542
  @{rail \<open>@{syntax method}\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   543
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   544
  \<^item> Terminal @{verbatim "'xyz'"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   545
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   546
  @{rail \<open>'xyz'\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   547
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   548
  \<^item> Terminal in keyword style @{verbatim "@'xyz'"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   549
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   550
  @{rail \<open>@'xyz'\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   551
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   552
  \<^item> Terminal via Isabelle antiquotation
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   553
  @{verbatim "@@{method rule}"}
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   554
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   555
  @{rail \<open>@@{method rule}\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   556
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   557
  \<^item> Concatenation @{verbatim "A B C"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   558
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   559
  @{rail \<open>A B C\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   560
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   561
  \<^item> Newline inside concatenation
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 54705
diff changeset
   562
  @{verbatim "A B C \<newline> D E F"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   563
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   564
  @{rail \<open>A B C \<newline> D E F\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   565
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   566
  \<^item> Variants @{verbatim "A | B | C"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   567
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   568
  @{rail \<open>A | B | C\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   569
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   570
  \<^item> Option @{verbatim "A ?"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   571
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   572
  @{rail \<open>A ?\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   573
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   574
  \<^item> Repetition @{verbatim "A *"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   575
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   576
  @{rail \<open>A *\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   577
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   578
  \<^item> Repetition with separator @{verbatim "A * sep"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   579
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   580
  @{rail \<open>A * sep\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   581
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   582
  \<^item> Strict repetition @{verbatim "A +"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   583
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   584
  @{rail \<open>A +\<close>}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   585
61421
e0825405d398 more symbols;
wenzelm
parents: 60286
diff changeset
   586
  \<^item> Strict repetition with separator @{verbatim "A + sep"}
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   587
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   588
  @{rail \<open>A + sep\<close>}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   589
\<close>
42658
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   590
8f5d5d71add0 some documentation of @{rail} antiquotation;
wenzelm
parents: 42651
diff changeset
   591
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   592
section \<open>Draft presentation\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   593
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   594
text \<open>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   595
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   596
    @{command_def "display_drafts"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   597
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   598
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   599
  @{rail \<open>
52549
802576856527 discontinued command 'print_drafts';
wenzelm
parents: 51960
diff changeset
   600
    @@{command display_drafts} (@{syntax name} +)
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   601
  \<close>}
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   602
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61491
diff changeset
   603
  \<^descr> @{command "display_drafts"}~\<open>paths\<close> performs simple output of a
52549
802576856527 discontinued command 'print_drafts';
wenzelm
parents: 51960
diff changeset
   604
  given list of raw source files. Only those symbols that do not require
802576856527 discontinued command 'print_drafts';
wenzelm
parents: 51960
diff changeset
   605
  additional {\LaTeX} packages are displayed properly, everything else is left
802576856527 discontinued command 'print_drafts';
wenzelm
parents: 51960
diff changeset
   606
  verbatim.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58593
diff changeset
   607
\<close>
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   608
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   609
end