| author | immler | 
| Thu, 27 Dec 2018 21:32:36 +0100 | |
| changeset 69513 | 42e08052dae8 | 
| parent 69041 | d57c460ba112 | 
| child 69551 | adb52af5ba55 | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 2 | ||
| 27037 | 3 | theory Outer_Syntax | 
| 63531 | 4 | imports Main Base | 
| 27037 | 5 | begin | 
| 6 | ||
| 58618 | 7 | chapter \<open>Outer syntax --- the theory language \label{ch:outer-syntax}\<close>
 | 
| 27037 | 8 | |
| 58618 | 9 | text \<open> | 
| 61662 | 10 | The rather generic framework of Isabelle/Isar syntax emerges from three main | 
| 11 | syntactic categories: \<^emph>\<open>commands\<close> of the top-level Isar engine (covering | |
| 12 | theory and proof elements), \<^emph>\<open>methods\<close> for general goal refinements | |
| 13 | (analogous to traditional ``tactics''), and \<^emph>\<open>attributes\<close> for operations on | |
| 14 | facts (within a certain context). Subsequently we give a reference of basic | |
| 15 | syntactic entities underlying Isabelle/Isar syntax in a bottom-up manner. | |
| 16 | Concrete theory and proof language elements will be introduced later on. | |
| 27037 | 17 | |
| 61421 | 18 | \<^medskip> | 
| 61662 | 19 | In order to get started with writing well-formed Isabelle/Isar documents, | 
| 20 | the most important aspect to be noted is the difference of \<^emph>\<open>inner\<close> versus | |
| 21 | \<^emph>\<open>outer\<close> syntax. Inner syntax is that of Isabelle types and terms of the | |
| 22 | logic, while outer syntax is that of Isabelle/Isar theory sources | |
| 23 | (specifications and proofs). As a general rule, inner syntax entities may | |
| 24 | occur only as \<^emph>\<open>atomic entities\<close> within outer syntax. For example, the | |
| 25 | string \<^verbatim>\<open>"x + y"\<close> and identifier \<^verbatim>\<open>z\<close> are legal term specifications within a | |
| 26 | theory, while \<^verbatim>\<open>x + y\<close> without quotes is not. | |
| 27037 | 27 | |
| 61662 | 28 | Printed theory documents usually omit quotes to gain readability (this is a | 
| 29 |   matter of {\LaTeX} macro setup, say via \<^verbatim>\<open>\isabellestyle\<close>, see also @{cite
 | |
| 30 | "isabelle-system"}). Experienced users of Isabelle/Isar may easily | |
| 31 | reconstruct the lost technical information, while mere readers need not care | |
| 32 | about quotes at all. | |
| 58618 | 33 | \<close> | 
| 27037 | 34 | |
| 35 | ||
| 58618 | 36 | section \<open>Commands\<close> | 
| 50213 | 37 | |
| 58618 | 38 | text \<open> | 
| 50213 | 39 |   \begin{matharray}{rcl}
 | 
| 61493 | 40 |     @{command_def "print_commands"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
 | 
| 41 |     @{command_def "help"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
 | |
| 50213 | 42 |   \end{matharray}
 | 
| 43 | ||
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 44 |   @{rail \<open>
 | 
| 50213 | 45 |     @@{command help} (@{syntax name} * )
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 46 | \<close>} | 
| 50213 | 47 | |
| 61439 | 48 |   \<^descr> @{command "print_commands"} prints all outer syntax keywords
 | 
| 50213 | 49 | and commands. | 
| 50 | ||
| 61493 | 51 |   \<^descr> @{command "help"}~\<open>pats\<close> retrieves outer syntax
 | 
| 50213 | 52 | commands according to the specified name patterns. | 
| 58618 | 53 | \<close> | 
| 50213 | 54 | |
| 55 | ||
| 58618 | 56 | subsubsection \<open>Examples\<close> | 
| 50213 | 57 | |
| 61662 | 58 | text \<open> | 
| 59 | Some common diagnostic commands are retrieved like this (according to usual | |
| 60 | naming conventions): | |
| 61 | \<close> | |
| 50213 | 62 | |
| 63 | help "print" | |
| 64 | help "find" | |
| 65 | ||
| 66 | ||
| 58618 | 67 | section \<open>Lexical matters \label{sec:outer-lex}\<close>
 | 
| 27037 | 68 | |
| 61662 | 69 | text \<open> | 
| 70 | The outer lexical syntax consists of three main categories of syntax tokens: | |
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 71 | |
| 61662 | 72 | \<^enum> \<^emph>\<open>major keywords\<close> --- the command names that are available | 
| 73 | in the present logic session; | |
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 74 | |
| 61662 | 75 | \<^enum> \<^emph>\<open>minor keywords\<close> --- additional literal tokens required | 
| 76 | by the syntax of commands; | |
| 27037 | 77 | |
| 61662 | 78 | \<^enum> \<^emph>\<open>named tokens\<close> --- various categories of identifiers etc. | 
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 79 | |
| 61662 | 80 | Major keywords and minor keywords are guaranteed to be disjoint. This helps | 
| 81 | user-interfaces to determine the overall structure of a theory text, without | |
| 82 | knowing the full details of command syntax. Internally, there is some | |
| 83 | additional information about the kind of major keywords, which approximates | |
| 84 | the command type (theory command, proof command etc.). | |
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 85 | |
| 61662 | 86 | Keywords override named tokens. For example, the presence of a command | 
| 87 | called \<^verbatim>\<open>term\<close> inhibits the identifier \<^verbatim>\<open>term\<close>, but the string \<^verbatim>\<open>"term"\<close> can | |
| 88 | be used instead. By convention, the outer syntax always allows quoted | |
| 89 | strings in addition to identifiers, wherever a named entity is expected. | |
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 90 | |
| 61662 | 91 | When tokenizing a given input sequence, the lexer repeatedly takes the | 
| 92 | longest prefix of the input that forms a valid token. Spaces, tabs, newlines | |
| 93 | and formfeeds between tokens serve as explicit separators. | |
| 28776 | 94 | |
| 61421 | 95 | \<^medskip> | 
| 96 | The categories for named tokens are defined once and for all as follows. | |
| 27037 | 97 | |
| 28776 | 98 |   \begin{center}
 | 
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 99 |   \begin{supertabular}{rcl}
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 100 |     @{syntax_def short_ident} & = & \<open>letter (subscript\<^sup>? quasiletter)\<^sup>*\<close> \\
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 101 |     @{syntax_def long_ident} & = & \<open>short_ident(\<close>\<^verbatim>\<open>.\<close>\<open>short_ident)\<^sup>+\<close> \\
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 102 |     @{syntax_def sym_ident} & = & \<open>sym\<^sup>+  |\<close>~~\<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>short_ident\<close>\<^verbatim>\<open>>\<close> \\
 | 
| 61493 | 103 |     @{syntax_def nat} & = & \<open>digit\<^sup>+\<close> \\
 | 
| 61503 | 104 |     @{syntax_def float} & = & @{syntax_ref nat}\<^verbatim>\<open>.\<close>@{syntax_ref nat}~~\<open>|\<close>~~\<^verbatim>\<open>-\<close>@{syntax_ref nat}\<^verbatim>\<open>.\<close>@{syntax_ref nat} \\
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 105 |     @{syntax_def term_var} & = & \<^verbatim>\<open>?\<close>\<open>short_ident  |\<close>~~\<^verbatim>\<open>?\<close>\<open>short_ident\<close>\<^verbatim>\<open>.\<close>\<open>nat\<close> \\
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 106 |     @{syntax_def type_ident} & = & \<^verbatim>\<open>'\<close>\<open>short_ident\<close> \\
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 107 |     @{syntax_def type_var} & = & \<^verbatim>\<open>?\<close>\<open>type_ident  |\<close>~~\<^verbatim>\<open>?\<close>\<open>type_ident\<close>\<^verbatim>\<open>.\<close>\<open>nat\<close> \\
 | 
| 61503 | 108 |     @{syntax_def string} & = & \<^verbatim>\<open>"\<close> \<open>\<dots>\<close> \<^verbatim>\<open>"\<close> \\
 | 
| 109 |     @{syntax_def altstring} & = & \<^verbatim>\<open>`\<close> \<open>\<dots>\<close> \<^verbatim>\<open>`\<close> \\
 | |
| 61493 | 110 |     @{syntax_def cartouche} & = & @{verbatim "\<open>"} \<open>\<dots>\<close> @{verbatim "\<close>"} \\
 | 
| 61503 | 111 |     @{syntax_def verbatim} & = & \<^verbatim>\<open>{*\<close> \<open>\<dots>\<close> \<^verbatim>\<open>*}\<close> \\[1ex]
 | 
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 112 | |
| 61503 | 113 | \<open>letter\<close> & = & \<open>latin |\<close>~~\<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>latin\<close>\<^verbatim>\<open>>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>latin latin\<close>\<^verbatim>\<open>>\<close>~~\<open>| greek |\<close> \\ | 
| 114 | \<open>subscript\<close> & = & \<^verbatim>\<open>\<^sub>\<close> \\ | |
| 115 | \<open>quasiletter\<close> & = & \<open>letter | digit |\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>'\<close> \\ | |
| 116 | \<open>latin\<close> & = & \<^verbatim>\<open>a\<close>~~\<open>| \<dots> |\<close>~~\<^verbatim>\<open>z\<close>~~\<open>|\<close>~~\<^verbatim>\<open>A\<close>~~\<open>| \<dots> |\<close>~~\<^verbatim>\<open>Z\<close> \\ | |
| 117 | \<open>digit\<close> & = & \<^verbatim>\<open>0\<close>~~\<open>| \<dots> |\<close>~~\<^verbatim>\<open>9\<close> \\ | |
| 118 | \<open>sym\<close> & = & \<^verbatim>\<open>!\<close>~~\<open>|\<close>~~\<^verbatim>\<open>#\<close>~~\<open>|\<close>~~\<^verbatim>\<open>$\<close>~~\<open>|\<close>~~\<^verbatim>\<open>%\<close>~~\<open>|\<close>~~\<^verbatim>\<open>&\<close>~~\<open>|\<close>~~\<^verbatim>\<open>*\<close>~~\<open>|\<close>~~\<^verbatim>\<open>+\<close>~~\<open>|\<close>~~\<^verbatim>\<open>-\<close>~~\<open>|\<close>~~\<^verbatim>\<open>/\<close>~~\<open>|\<close> \\ | |
| 119 | & & \<^verbatim>\<open><\<close>~~\<open>|\<close>~~\<^verbatim>\<open>=\<close>~~\<open>|\<close>~~\<^verbatim>\<open>>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>?\<close>~~\<open>|\<close>~~\<^verbatim>\<open>@\<close>~~\<open>|\<close>~~\<^verbatim>\<open>^\<close>~~\<open>|\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>|\<close>~~\<open>|\<close>~~\<^verbatim>\<open>~\<close> \\ | |
| 120 | \<open>greek\<close> & = & \<^verbatim>\<open>\<alpha>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<beta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<gamma>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<delta>\<close>~~\<open>|\<close> \\ | |
| 121 | & & \<^verbatim>\<open>\<epsilon>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<zeta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<eta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<theta>\<close>~~\<open>|\<close> \\ | |
| 122 | & & \<^verbatim>\<open>\<iota>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<kappa>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<mu>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<nu>\<close>~~\<open>|\<close> \\ | |
| 123 | & & \<^verbatim>\<open>\<xi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<pi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<rho>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<sigma>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<tau>\<close>~~\<open>|\<close> \\ | |
| 124 | & & \<^verbatim>\<open>\<upsilon>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<phi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<chi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<psi>\<close>~~\<open>|\<close> \\ | |
| 125 | & & \<^verbatim>\<open>\<omega>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Gamma>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Delta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Theta>\<close>~~\<open>|\<close> \\ | |
| 126 | & & \<^verbatim>\<open>\<Lambda>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Xi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Pi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Sigma>\<close>~~\<open>|\<close> \\ | |
| 127 | & & \<^verbatim>\<open>\<Upsilon>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Phi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Psi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Omega>\<close> \\ | |
| 28775 
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
 wenzelm parents: 
28774diff
changeset | 128 |   \end{supertabular}
 | 
| 28776 | 129 |   \end{center}
 | 
| 27037 | 130 | |
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 131 |   A @{syntax_ref term_var} or @{syntax_ref type_var} describes an unknown,
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 132 |   which is internally a pair of base name and index (ML type @{ML_type
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 133 | indexname}). These components are either separated by a dot as in \<open>?x.1\<close> or | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 134 | \<open>?x7.3\<close> or run together as in \<open>?x1\<close>. The latter form is possible if the base | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 135 | name does not end with digits. If the index is 0, it may be dropped | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 136 | altogether: \<open>?x\<close> and \<open>?x0\<close> and \<open>?x.0\<close> all refer to the same unknown, with | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 137 | basename \<open>x\<close> and index 0. | 
| 28778 | 138 | |
| 139 |   The syntax of @{syntax_ref string} admits any characters, including
 | |
| 61662 | 140 | newlines; ``\<^verbatim>\<open>"\<close>'' (double-quote) and ``\<^verbatim>\<open>\\<close>'' (backslash) need to be | 
| 141 | escaped by a backslash; arbitrary character codes may be specified as | |
| 142 | ``\<^verbatim>\<open>\\<close>\<open>ddd\<close>'', with three decimal digits. Alternative strings according to | |
| 143 |   @{syntax_ref altstring} are analogous, using single back-quotes instead.
 | |
| 28778 | 144 | |
| 58725 | 145 |   The body of @{syntax_ref verbatim} may consist of any text not containing
 | 
| 61662 | 146 | ``\<^verbatim>\<open>*}\<close>''; this allows to include quotes without further escapes, but there | 
| 147 | is no way to escape ``\<^verbatim>\<open>*}\<close>''. Cartouches do not have this limitation. | |
| 28778 | 148 | |
| 61662 | 149 |   A @{syntax_ref cartouche} consists of arbitrary text, with properly balanced
 | 
| 150 |   blocks of ``@{verbatim "\<open>"}~\<open>\<dots>\<close>~@{verbatim "\<close>"}''. Note that the rendering
 | |
| 151 | of cartouche delimiters is usually like this: ``\<open>\<open> \<dots> \<close>\<close>''. | |
| 55033 | 152 | |
| 69041 | 153 | Source comments take the form \<^verbatim>\<open>(*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*)\<close> and may be nested: the text is | 
| 154 | removed after lexical analysis of the input and thus not suitable for | |
| 155 | documentation. The Isar syntax also provides proper \<^emph>\<open>document comments\<close> | |
| 156 |   that are considered as part of the text (see \secref{sec:comments}).
 | |
| 27037 | 157 | |
| 61662 | 158 | Common mathematical symbols such as \<open>\<forall>\<close> are represented in Isabelle as \<^verbatim>\<open>\<forall>\<close>. | 
| 159 | There are infinitely many Isabelle symbols like this, although proper | |
| 160 |   presentation is left to front-end tools such as {\LaTeX} or Isabelle/jEdit.
 | |
| 161 | A list of predefined Isabelle symbols that work well with these tools is | |
| 162 |   given in \appref{app:symbols}. Note that \<^verbatim>\<open>\<lambda>\<close> does not belong to the
 | |
| 163 | \<open>letter\<close> category, since it is already used differently in the Pure term | |
| 164 | language. | |
| 165 | \<close> | |
| 27037 | 166 | |
| 167 | ||
| 58618 | 168 | section \<open>Common syntax entities\<close> | 
| 27037 | 169 | |
| 58618 | 170 | text \<open> | 
| 61662 | 171 | We now introduce several basic syntactic entities, such as names, terms, and | 
| 172 | theorem specifications, which are factored out of the actual Isar language | |
| 173 | elements to be described later. | |
| 58618 | 174 | \<close> | 
| 27037 | 175 | |
| 176 | ||
| 58618 | 177 | subsection \<open>Names\<close> | 
| 27037 | 178 | |
| 61662 | 179 | text \<open> | 
| 180 |   Entity @{syntax name} usually refers to any name of types, constants,
 | |
| 62969 | 181 | theorems etc.\ Quoted strings provide an escape for non-identifier names or | 
| 182 | those ruled out by outer syntax keywords (e.g.\ quoted \<^verbatim>\<open>"let"\<close>). | |
| 27037 | 183 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 184 |   @{rail \<open>
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 185 |     @{syntax_def name}: @{syntax short_ident} | @{syntax long_ident} |
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 186 |       @{syntax sym_ident} | @{syntax nat} | @{syntax string}
 | 
| 27037 | 187 | ; | 
| 60131 | 188 |     @{syntax_def par_name}: '(' @{syntax name} ')'
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 189 | \<close>} | 
| 66916 | 190 | |
| 191 |   A @{syntax_def system_name} is like @{syntax name}, but it excludes
 | |
| 192 | white-space characters and needs to conform to file-name notation. | |
| 58618 | 193 | \<close> | 
| 40296 | 194 | |
| 195 | ||
| 58618 | 196 | subsection \<open>Numbers\<close> | 
| 40296 | 197 | |
| 61662 | 198 | text \<open> | 
| 199 |   The outer lexical syntax (\secref{sec:outer-lex}) admits natural numbers and
 | |
| 200 |   floating point numbers. These are combined as @{syntax int} and @{syntax
 | |
| 201 | real} as follows. | |
| 40296 | 202 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 203 |   @{rail \<open>
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 204 |     @{syntax_def int}: @{syntax nat} | '-' @{syntax nat}
 | 
| 27037 | 205 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 206 |     @{syntax_def real}: @{syntax float} | @{syntax int}
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 207 | \<close>} | 
| 40296 | 208 | |
| 61662 | 209 |   Note that there is an overlap with the category @{syntax name}, which also
 | 
| 210 |   includes @{syntax nat}.
 | |
| 58618 | 211 | \<close> | 
| 27037 | 212 | |
| 213 | ||
| 63120 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 214 | subsection \<open>Embedded content\<close> | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 215 | |
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 216 | text \<open> | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 217 |   Entity @{syntax embedded} refers to content of other languages: cartouches
 | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 218 | allow arbitrary nesting of sub-languages that respect the recursive | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 219 | balancing of cartouche delimiters. Quoted strings are possible as well, but | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 220 | require escaped quotes when nested. As a shortcut, tokens that appear as | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 221 | plain identifiers in the outer language may be used as inner language | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 222 | content without delimiters. | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 223 | |
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 224 |   @{rail \<open>
 | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 225 |     @{syntax_def embedded}: @{syntax cartouche} | @{syntax string} |
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 226 |       @{syntax short_ident} | @{syntax long_ident} | @{syntax sym_ident} |
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 227 |       @{syntax term_var} | @{syntax type_ident} | @{syntax type_var} | @{syntax nat}
 | 
| 63120 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 228 | \<close>} | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 229 | \<close> | 
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 230 | |
| 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
62969diff
changeset | 231 | |
| 67448 | 232 | subsection \<open>Document text\<close> | 
| 27037 | 233 | |
| 61579 | 234 | text \<open> | 
| 67448 | 235 |   A chunk of document @{syntax text} is usually given as @{syntax cartouche}
 | 
| 236 |   \<open>\<open>\<dots>\<close>\<close> or @{syntax verbatim}, i.e.\ enclosed in \<^verbatim>\<open>{*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*}\<close>. For
 | |
| 237 |   convenience, any of the smaller text unit that conforms to @{syntax name} is
 | |
| 238 | admitted as well. | |
| 27037 | 239 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 240 |   @{rail \<open>
 | 
| 63139 | 241 |     @{syntax_def text}: @{syntax embedded} | @{syntax verbatim}
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 242 | \<close>} | 
| 67448 | 243 | |
| 244 | Typical uses are document markup commands, like \<^theory_text>\<open>chapter\<close>, \<^theory_text>\<open>section\<close> etc. | |
| 245 |   (\secref{sec:markup}).
 | |
| 246 | \<close> | |
| 247 | ||
| 248 | ||
| 249 | subsection \<open>Document comments \label{sec:comments}\<close>
 | |
| 250 | ||
| 251 | text \<open> | |
| 252 | Formal comments are an integral part of the document, but are logically void | |
| 253 | and removed from the resulting theory or term content. The output of | |
| 254 |   document preparation (\chref{ch:document-prep}) supports various styles,
 | |
| 255 | according to the following kinds of comments. | |
| 256 | ||
| 67476 | 257 | \<^item> Marginal comment of the form \<^verbatim>\<open>\<comment>\<close>~\<open>\<open>text\<close>\<close> or \<open>\<comment>\<close>~\<open>\<open>text\<close>\<close>, usually with | 
| 258 | a single space between the comment symbol and the argument cartouche. The | |
| 67448 | 259 | given argument is typeset as regular text, with formal antiquotations | 
| 260 |     (\secref{sec:antiq}).
 | |
| 261 | ||
| 262 | \<^item> Canceled text of the form \<^verbatim>\<open>\<^cancel>\<close>\<open>\<open>text\<close>\<close> (no white space between the | |
| 263 | control symbol and the argument cartouche). The argument is typeset as | |
| 264 | formal Isabelle source and overlaid with a ``strike-through'' pattern, | |
| 265 | e.g. \<^theory_text>\<open>\<^cancel>\<open>bad\<close>\<close>. | |
| 266 | ||
| 267 |     \<^item> Raw {\LaTeX} source of the form \<^verbatim>\<open>\<^latex>\<close>\<open>\<open>argument\<close>\<close> (no white space
 | |
| 268 | between the control symbol and the argument cartouche). This allows to | |
| 269 |     augment the generated {\TeX} source arbitrarily, without any sanity
 | |
| 270 | checks! | |
| 271 | ||
| 272 | These formal comments work uniformly in outer syntax, inner syntax (term | |
| 273 | language), Isabelle/ML, and some other embedded languages of Isabelle. | |
| 58618 | 274 | \<close> | 
| 27037 | 275 | |
| 276 | ||
| 58618 | 277 | subsection \<open>Type classes, sorts and arities\<close> | 
| 27037 | 278 | |
| 58618 | 279 | text \<open> | 
| 61662 | 280 | Classes are specified by plain names. Sorts have a very simple inner syntax, | 
| 281 |   which is either a single class name \<open>c\<close> or a list \<open>{c\<^sub>1, \<dots>, c\<^sub>n}\<close> referring
 | |
| 282 | to the intersection of these classes. The syntax of type arities is given | |
| 27037 | 283 | directly at the outer level. | 
| 284 | ||
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 285 |   @{rail \<open>
 | 
| 62969 | 286 |     @{syntax_def classdecl}: @{syntax name} (('<' | '\<subseteq>') (@{syntax name} + ','))?
 | 
| 27037 | 287 | ; | 
| 63140 | 288 |     @{syntax_def sort}: @{syntax embedded}
 | 
| 27037 | 289 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 290 |     @{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort}
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 291 | \<close>} | 
| 58618 | 292 | \<close> | 
| 27037 | 293 | |
| 294 | ||
| 58618 | 295 | subsection \<open>Types and terms \label{sec:types-terms}\<close>
 | 
| 27037 | 296 | |
| 58618 | 297 | text \<open> | 
| 61662 | 298 | The actual inner Isabelle syntax, that of types and terms of the logic, is | 
| 299 | far too sophisticated in order to be modelled explicitly at the outer theory | |
| 300 | level. Basically, any such entity has to be quoted to turn it into a single | |
| 301 | token (the parsing and type-checking is performed internally later). For | |
| 302 | convenience, a slightly more liberal convention is adopted: quotes may be | |
| 303 | omitted for any type or term that is already atomic at the outer level. For | |
| 304 | example, one may just write \<^verbatim>\<open>x\<close> instead of quoted \<^verbatim>\<open>"x"\<close>. Note that | |
| 305 | symbolic identifiers (e.g.\ \<^verbatim>\<open>++\<close> or \<open>\<forall>\<close> are available as well, provided | |
| 306 | these have not been superseded by commands or other keywords already (such | |
| 307 | as \<^verbatim>\<open>=\<close> or \<^verbatim>\<open>+\<close>). | |
| 27037 | 308 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 309 |   @{rail \<open>
 | 
| 63137 
9553f11d67c4
simplified syntax: Parse.term corresponds to Args.term etc.;
 wenzelm parents: 
63120diff
changeset | 310 |     @{syntax_def type}: @{syntax embedded}
 | 
| 27037 | 311 | ; | 
| 63137 
9553f11d67c4
simplified syntax: Parse.term corresponds to Args.term etc.;
 wenzelm parents: 
63120diff
changeset | 312 |     @{syntax_def term}: @{syntax embedded}
 | 
| 27037 | 313 | ; | 
| 63137 
9553f11d67c4
simplified syntax: Parse.term corresponds to Args.term etc.;
 wenzelm parents: 
63120diff
changeset | 314 |     @{syntax_def prop}: @{syntax embedded}
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 315 | \<close>} | 
| 27037 | 316 | |
| 59853 | 317 | Positional instantiations are specified as a sequence of terms, or the | 
| 61493 | 318 | placeholder ``\<open>_\<close>'' (underscore), which means to skip a position. | 
| 27037 | 319 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 320 |   @{rail \<open>
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 321 |     @{syntax_def inst}: '_' | @{syntax term}
 | 
| 27037 | 322 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 323 |     @{syntax_def insts}: (@{syntax inst} *)
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 324 | \<close>} | 
| 27037 | 325 | |
| 61662 | 326 | Named instantiations are specified as pairs of assignments \<open>v = t\<close>, which | 
| 327 | refer to schematic variables in some theorem that is instantiated. Both type | |
| 328 | and terms instantiations are admitted, and distinguished by the usual syntax | |
| 329 | of variable names. | |
| 59853 | 330 | |
| 331 |   @{rail \<open>
 | |
| 332 |     @{syntax_def named_inst}: variable '=' (type | term)
 | |
| 333 | ; | |
| 334 |     @{syntax_def named_insts}: (named_inst @'and' +)
 | |
| 335 | ; | |
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 336 |     variable: @{syntax name} | @{syntax term_var} | @{syntax type_ident} | @{syntax type_var}
 | 
| 59853 | 337 | \<close>} | 
| 338 | ||
| 61662 | 339 |   Type declarations and definitions usually refer to @{syntax typespec} on the
 | 
| 340 | left-hand side. This models basic type constructor application at the outer | |
| 341 | syntax level. Note that only plain postfix notation is available here, but | |
| 342 | no infixes. | |
| 27037 | 343 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 344 |   @{rail \<open>
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 345 |     @{syntax_def typespec}:
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 346 |       (() | @{syntax type_ident} | '(' ( @{syntax type_ident} + ',' ) ')') @{syntax name}
 | 
| 27037 | 347 | ; | 
| 42705 | 348 |     @{syntax_def typespec_sorts}:
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 349 |       (() | (@{syntax type_ident} ('::' @{syntax sort})?) |
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 350 |         '(' ( (@{syntax type_ident} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name}
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 351 | \<close>} | 
| 58618 | 352 | \<close> | 
| 27037 | 353 | |
| 354 | ||
| 58618 | 355 | subsection \<open>Term patterns and declarations \label{sec:term-decls}\<close>
 | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 356 | |
| 61662 | 357 | text \<open> | 
| 358 | Wherever explicit propositions (or term fragments) occur in a proof text, | |
| 359 | casual binding of schematic term variables may be given specified via | |
| 61663 | 360 |   patterns of the form ``\<^theory_text>\<open>(is p\<^sub>1 \<dots> p\<^sub>n)\<close>''. This works both for @{syntax
 | 
| 61662 | 361 |   term} and @{syntax prop}.
 | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 362 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 363 |   @{rail \<open>
 | 
| 42705 | 364 |     @{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')'
 | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 365 | ; | 
| 42705 | 366 |     @{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')'
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 367 | \<close>} | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 368 | |
| 61421 | 369 | \<^medskip> | 
| 61662 | 370 | Declarations of local variables \<open>x :: \<tau>\<close> and logical propositions \<open>a : \<phi>\<close> | 
| 371 | represent different views on the same principle of introducing a local | |
| 372 |   scope. In practice, one may usually omit the typing of @{syntax vars} (due
 | |
| 373 | to type-inference), and the naming of propositions (due to implicit | |
| 374 | references of current facts). In any case, Isar proof elements usually admit | |
| 375 | to introduce multiple such items simultaneously. | |
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 376 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 377 |   @{rail \<open>
 | 
| 63285 | 378 |     @{syntax_def vars}:
 | 
| 379 |       (((@{syntax name} +) ('::' @{syntax type})? |
 | |
| 380 |         @{syntax name} ('::' @{syntax type})? @{syntax mixfix}) + @'and')
 | |
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 381 | ; | 
| 42705 | 382 |     @{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +)
 | 
| 61658 | 383 | ; | 
| 384 |     @{syntax_def props'}: (@{syntax prop} @{syntax prop_pat}? +)
 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 385 | \<close>} | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 386 | |
| 61662 | 387 | The treatment of multiple declarations corresponds to the complementary | 
| 388 |   focus of @{syntax vars} versus @{syntax props}. In ``\<open>x\<^sub>1 \<dots> x\<^sub>n :: \<tau>\<close>'' the
 | |
| 389 | typing refers to all variables, while in \<open>a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n\<close> the naming refers to | |
| 390 |   all propositions collectively. Isar language elements that refer to @{syntax
 | |
| 391 |   vars} or @{syntax props} typically admit separate typings or namings via
 | |
| 392 |   another level of iteration, with explicit @{keyword_ref "and"} separators;
 | |
| 393 |   e.g.\ see @{command "fix"} and @{command "assume"} in
 | |
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 394 |   \secref{sec:proof-context}.
 | 
| 61662 | 395 | \<close> | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 396 | |
| 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
28753diff
changeset | 397 | |
| 58618 | 398 | subsection \<open>Attributes and theorems \label{sec:syn-att}\<close>
 | 
| 27037 | 399 | |
| 61662 | 400 | text \<open> | 
| 401 | Attributes have their own ``semi-inner'' syntax, in the sense that input | |
| 402 |   conforming to @{syntax args} below is parsed by the attribute a second time.
 | |
| 403 | The attribute argument specifications may be any sequence of atomic entities | |
| 404 | (identifiers, strings etc.), or properly bracketed argument lists. Below | |
| 405 |   @{syntax atom} refers to any atomic entity, including any @{syntax keyword}
 | |
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 406 |   conforming to @{syntax sym_ident}.
 | 
| 27037 | 407 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 408 |   @{rail \<open>
 | 
| 63138 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 409 |     @{syntax_def atom}: @{syntax name} | @{syntax type_ident} |
 | 
| 
70f4d67235a0
clarified syntax category names according to Isabelle/ML/Scala;
 wenzelm parents: 
63137diff
changeset | 410 |       @{syntax type_var} | @{syntax term_var} | @{syntax nat} | @{syntax float} |
 | 
| 55045 
99056d23e05b
cartouche within nested args (attributes, methods, etc.);
 wenzelm parents: 
55033diff
changeset | 411 |       @{syntax keyword} | @{syntax cartouche}
 | 
| 27037 | 412 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 413 |     arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']'
 | 
| 27037 | 414 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 415 |     @{syntax_def args}: arg *
 | 
| 27037 | 416 | ; | 
| 62969 | 417 |     @{syntax_def attributes}: '[' (@{syntax name} @{syntax args} * ',') ']'
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 418 | \<close>} | 
| 27037 | 419 | |
| 61662 | 420 |   Theorem specifications come in several flavors: @{syntax axmdecl} and
 | 
| 421 |   @{syntax thmdecl} usually refer to axioms, assumptions or results of goal
 | |
| 422 |   statements, while @{syntax thmdef} collects lists of existing theorems.
 | |
| 62969 | 423 |   Existing theorems are given by @{syntax thm} and @{syntax thms}, the
 | 
| 61662 | 424 | former requires an actual singleton result. | 
| 27037 | 425 | |
| 426 | There are three forms of theorem references: | |
| 60674 | 427 | |
| 61662 | 428 | \<^enum> named facts \<open>a\<close>, | 
| 27037 | 429 | |
| 61662 | 430 | \<^enum> selections from named facts \<open>a(i)\<close> or \<open>a(j - k)\<close>, | 
| 27037 | 431 | |
| 61662 | 432 |     \<^enum> literal fact propositions using token syntax @{syntax_ref altstring}
 | 
| 433 |     \<^verbatim>\<open>`\<close>\<open>\<phi>\<close>\<^verbatim>\<open>`\<close> or @{syntax_ref cartouche}
 | |
| 434 |     \<open>\<open>\<phi>\<close>\<close> (see also method @{method_ref fact}).
 | |
| 27037 | 435 | |
| 61662 | 436 | Any kind of theorem specification may include lists of attributes both on | 
| 437 | the left and right hand sides; attributes are applied to any immediately | |
| 438 | preceding fact. If names are omitted, the theorems are not stored within the | |
| 439 | theorem database of the theory or proof context, but any given attributes | |
| 440 | are applied nonetheless. | |
| 27037 | 441 | |
| 61662 | 442 | An extra pair of brackets around attributes (like ``\<open>[[simproc a]]\<close>'') | 
| 443 | abbreviates a theorem reference involving an internal dummy fact, which will | |
| 444 | be ignored later on. So only the effect of the attribute on the background | |
| 445 | context will persist. This form of in-place declarations is particularly | |
| 446 |   useful with commands like @{command "declare"} and @{command "using"}.
 | |
| 27037 | 447 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 448 |   @{rail \<open>
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 449 |     @{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':'
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 450 | ; | 
| 60631 | 451 |     @{syntax_def thmbind}:
 | 
| 452 |       @{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes}
 | |
| 453 | ; | |
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 454 |     @{syntax_def thmdecl}: thmbind ':'
 | 
| 27037 | 455 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 456 |     @{syntax_def thmdef}: thmbind '='
 | 
| 27037 | 457 | ; | 
| 62969 | 458 |     @{syntax_def thm}:
 | 
| 459 |       (@{syntax name} selection? | @{syntax altstring} | @{syntax cartouche})
 | |
| 56499 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56451diff
changeset | 460 |         @{syntax attributes}? |
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 461 |       '[' @{syntax attributes} ']'
 | 
| 27037 | 462 | ; | 
| 62969 | 463 |     @{syntax_def thms}: @{syntax thm} +
 | 
| 27037 | 464 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40296diff
changeset | 465 |     selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')'
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55045diff
changeset | 466 | \<close>} | 
| 58618 | 467 | \<close> | 
| 27037 | 468 | |
| 60674 | 469 | |
| 63182 | 470 | subsection \<open>Structured specifications\<close> | 
| 471 | ||
| 472 | text \<open> | |
| 473 | Structured specifications use propositions with explicit notation for the | |
| 474 | ``eigen-context'' to describe rule structure: \<open>\<And>x. A x \<Longrightarrow> \<dots> \<Longrightarrow> B x\<close> is | |
| 475 | expressed as \<^theory_text>\<open>B x if A x and \<dots> for x\<close>. It is also possible to use dummy | |
| 476 | terms ``\<open>_\<close>'' (underscore) to refer to locally fixed variables anonymously. | |
| 477 | ||
| 478 | Multiple specifications are delimited by ``\<open>|\<close>'' to emphasize separate | |
| 479 | cases: each with its own scope of inferred types for free variables. | |
| 480 | ||
| 481 | ||
| 482 |   @{rail \<open>
 | |
| 63285 | 483 |     @{syntax_def for_fixes}: (@'for' @{syntax vars})?
 | 
| 484 | ; | |
| 63182 | 485 |     @{syntax_def multi_specs}: (@{syntax structured_spec} + '|')
 | 
| 486 | ; | |
| 487 |     @{syntax_def structured_spec}:
 | |
| 488 |       @{syntax thmdecl}? @{syntax prop} @{syntax spec_prems} @{syntax for_fixes}
 | |
| 489 | ; | |
| 490 |     @{syntax_def spec_prems}: (@'if' ((@{syntax prop}+) + @'and'))?
 | |
| 63183 | 491 | ; | 
| 63285 | 492 |     @{syntax_def specification}: @{syntax vars} @'where' @{syntax multi_specs}
 | 
| 63182 | 493 | \<close>} | 
| 494 | \<close> | |
| 495 | ||
| 496 | ||
| 60674 | 497 | section \<open>Diagnostic commands\<close> | 
| 498 | ||
| 499 | text \<open> | |
| 500 |   \begin{matharray}{rcl}
 | |
| 61493 | 501 |     @{command_def "print_theory"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | 
| 502 |     @{command_def "print_definitions"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 503 |     @{command_def "print_methods"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 504 |     @{command_def "print_attributes"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 505 |     @{command_def "print_theorems"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 506 |     @{command_def "find_theorems"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 507 |     @{command_def "find_consts"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 508 |     @{command_def "thm_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 509 |     @{command_def "unused_thms"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 510 |     @{command_def "print_facts"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 511 |     @{command_def "print_term_bindings"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 60674 | 512 |   \end{matharray}
 | 
| 513 | ||
| 514 |   @{rail \<open>
 | |
| 515 |     (@@{command print_theory} |
 | |
| 61252 | 516 |       @@{command print_definitions} |
 | 
| 60674 | 517 |       @@{command print_methods} |
 | 
| 518 |       @@{command print_attributes} |
 | |
| 519 |       @@{command print_theorems} |
 | |
| 520 |       @@{command print_facts}) ('!'?)
 | |
| 521 | ; | |
| 522 |     @@{command find_theorems} ('(' @{syntax nat}? 'with_dups'? ')')? \<newline> (thm_criterion*)
 | |
| 523 | ; | |
| 62969 | 524 |     thm_criterion: ('-'?) ('name' ':' @{syntax name} | 'intro' | 'elim' | 'dest' |
 | 
| 60674 | 525 |       'solves' | 'simp' ':' @{syntax term} | @{syntax term})
 | 
| 526 | ; | |
| 527 |     @@{command find_consts} (const_criterion*)
 | |
| 528 | ; | |
| 529 |     const_criterion: ('-'?)
 | |
| 62969 | 530 |       ('name' ':' @{syntax name} | 'strict' ':' @{syntax type} | @{syntax type})
 | 
| 60674 | 531 | ; | 
| 532 |     @@{command thm_deps} @{syntax thmrefs}
 | |
| 533 | ; | |
| 534 |     @@{command unused_thms} ((@{syntax name} +) '-' (@{syntax name} * ))?
 | |
| 535 | \<close>} | |
| 536 | ||
| 61662 | 537 | These commands print certain parts of the theory and proof context. Note | 
| 538 | that there are some further ones available, such as for the set of rules | |
| 539 | declared for simplifications. | |
| 60674 | 540 | |
| 61439 | 541 |   \<^descr> @{command "print_theory"} prints the main logical content of the
 | 
| 61493 | 542 | background theory; the ``\<open>!\<close>'' option indicates extra verbosity. | 
| 60674 | 543 | |
| 61439 | 544 |   \<^descr> @{command "print_definitions"} prints dependencies of definitional
 | 
| 61252 | 545 | specifications within the background theory, which may be constants | 
| 62172 
7eaeae127955
updated section on "Overloaded constant definitions";
 wenzelm parents: 
61663diff
changeset | 546 |   (\secref{sec:term-definitions}, \secref{sec:overloading}) or types
 | 
| 
7eaeae127955
updated section on "Overloaded constant definitions";
 wenzelm parents: 
61663diff
changeset | 547 |   (\secref{sec:types-pure}, \secref{sec:hol-typedef}); the ``\<open>!\<close>'' option
 | 
| 
7eaeae127955
updated section on "Overloaded constant definitions";
 wenzelm parents: 
61663diff
changeset | 548 | indicates extra verbosity. | 
| 61252 | 549 | |
| 61439 | 550 |   \<^descr> @{command "print_methods"} prints all proof methods available in the
 | 
| 61662 | 551 | current theory context; the ``\<open>!\<close>'' option indicates extra verbosity. | 
| 60674 | 552 | |
| 61439 | 553 |   \<^descr> @{command "print_attributes"} prints all attributes available in the
 | 
| 61662 | 554 | current theory context; the ``\<open>!\<close>'' option indicates extra verbosity. | 
| 60674 | 555 | |
| 61439 | 556 |   \<^descr> @{command "print_theorems"} prints theorems of the background theory
 | 
| 61662 | 557 | resulting from the last command; the ``\<open>!\<close>'' option indicates extra | 
| 558 | verbosity. | |
| 60674 | 559 | |
| 61662 | 560 |   \<^descr> @{command "print_facts"} prints all local facts of the current context,
 | 
| 561 | both named and unnamed ones; the ``\<open>!\<close>'' option indicates extra verbosity. | |
| 60674 | 562 | |
| 61662 | 563 |   \<^descr> @{command "print_term_bindings"} prints all term bindings that are present
 | 
| 564 | in the context. | |
| 60674 | 565 | |
| 61662 | 566 |   \<^descr> @{command "find_theorems"}~\<open>criteria\<close> retrieves facts from the theory or
 | 
| 567 | proof context matching all of given search criteria. The criterion \<open>name: p\<close> | |
| 568 | selects all theorems whose fully qualified name matches pattern \<open>p\<close>, which | |
| 569 | may contain ``\<open>*\<close>'' wildcards. The criteria \<open>intro\<close>, \<open>elim\<close>, and \<open>dest\<close> | |
| 570 | select theorems that match the current goal as introduction, elimination or | |
| 571 | destruction rules, respectively. The criterion \<open>solves\<close> returns all rules | |
| 572 | that would directly solve the current goal. The criterion \<open>simp: t\<close> selects | |
| 573 | all rewrite rules whose left-hand side matches the given term. The criterion | |
| 574 | term \<open>t\<close> selects all theorems that contain the pattern \<open>t\<close> -- as usual, | |
| 575 | patterns may contain occurrences of the dummy ``\<open>_\<close>'', schematic variables, | |
| 576 | and type constraints. | |
| 60674 | 577 | |
| 61662 | 578 | Criteria can be preceded by ``\<open>-\<close>'' to select theorems that do \<^emph>\<open>not\<close> match. | 
| 579 | Note that giving the empty list of criteria yields \<^emph>\<open>all\<close> currently known | |
| 580 | facts. An optional limit for the number of printed facts may be given; the | |
| 581 | default is 40. By default, duplicates are removed from the search result. | |
| 582 | Use \<open>with_dups\<close> to display duplicates. | |
| 60674 | 583 | |
| 61662 | 584 |   \<^descr> @{command "find_consts"}~\<open>criteria\<close> prints all constants whose type meets
 | 
| 585 | all of the given criteria. The criterion \<open>strict: ty\<close> is met by any type | |
| 586 | that matches the type pattern \<open>ty\<close>. Patterns may contain both the dummy type | |
| 587 | ``\<open>_\<close>'' and sort constraints. The criterion \<open>ty\<close> is similar, but it also | |
| 588 | matches against subtypes. The criterion \<open>name: p\<close> and the prefix ``\<open>-\<close>'' | |
| 589 |   function as described for @{command "find_theorems"}.
 | |
| 60674 | 590 | |
| 61662 | 591 |   \<^descr> @{command "thm_deps"}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> visualizes dependencies of facts, using
 | 
| 592 |   Isabelle's graph browser tool (see also @{cite "isabelle-system"}).
 | |
| 60674 | 593 | |
| 61662 | 594 |   \<^descr> @{command "unused_thms"}~\<open>A\<^sub>1 \<dots> A\<^sub>m - B\<^sub>1 \<dots> B\<^sub>n\<close> displays all theorems
 | 
| 595 | that are proved in theories \<open>B\<^sub>1 \<dots> B\<^sub>n\<close> or their parents but not in \<open>A\<^sub>1 \<dots> | |
| 596 | A\<^sub>m\<close> or their parents and that are never used. If \<open>n\<close> is \<open>0\<close>, the end of the | |
| 597 | range of theories defaults to the current theory. If no range is specified, | |
| 60674 | 598 | only the unused theorems in the current theory are displayed. | 
| 599 | \<close> | |
| 600 | ||
| 27037 | 601 | end |