| author | wenzelm | 
| Thu, 13 Nov 2008 21:49:46 +0100 | |
| changeset 28763 | b5e6122ff575 | 
| parent 28762 | f5d79aeffd81 | 
| child 28765 | da8f6f4a74be | 
| permissions | -rw-r--r-- | 
| 28762 | 1 | (* $Id$ *) | 
| 2 | ||
| 3 | theory Inner_Syntax | |
| 4 | imports Main | |
| 5 | begin | |
| 6 | ||
| 7 | chapter {* Inner syntax --- the term language *}
 | |
| 8 | ||
| 9 | section {* Printing logical entities *}
 | |
| 10 | ||
| 11 | subsection {* Diagnostic commands *}
 | |
| 12 | ||
| 13 | text {*
 | |
| 14 |   \begin{matharray}{rcl}
 | |
| 15 |     @{command_def "pr"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
 | |
| 16 |     @{command_def "thm"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | |
| 17 |     @{command_def "term"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | |
| 18 |     @{command_def "prop"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | |
| 19 |     @{command_def "typ"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | |
| 20 |     @{command_def "prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | |
| 21 |     @{command_def "full_prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | |
| 22 |   \end{matharray}
 | |
| 23 | ||
| 24 | These diagnostic commands assist interactive development by printing | |
| 25 | internal logical entities in a human-readable fashion. | |
| 26 | ||
| 27 |   \begin{rail}
 | |
| 28 |     'pr' modes? nat? (',' nat)?
 | |
| 29 | ; | |
| 30 | 'thm' modes? thmrefs | |
| 31 | ; | |
| 32 | 'term' modes? term | |
| 33 | ; | |
| 34 | 'prop' modes? prop | |
| 35 | ; | |
| 36 | 'typ' modes? type | |
| 37 | ; | |
| 38 | 'prf' modes? thmrefs? | |
| 39 | ; | |
| 40 | 'full\_prf' modes? thmrefs? | |
| 41 | ; | |
| 42 | ||
| 43 |     modes: '(' (name + ) ')'
 | |
| 44 | ; | |
| 45 |   \end{rail}
 | |
| 46 | ||
| 47 |   \begin{description}
 | |
| 48 | ||
| 49 |   \item @{command "pr"}~@{text "goals, prems"} prints the current
 | |
| 50 | proof state (if present), including the proof context, current facts | |
| 51 | and goals. The optional limit arguments affect the number of goals | |
| 52 | and premises to be displayed, which is initially 10 for both. | |
| 53 | Omitting limit values leaves the current setting unchanged. | |
| 54 | ||
| 55 |   \item @{command "thm"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} retrieves
 | |
| 56 | theorems from the current theory or proof context. Note that any | |
| 57 | attributes included in the theorem specifications are applied to a | |
| 58 | temporary context derived from the current theory or proof; the | |
| 59 |   result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1,
 | |
| 60 | \<dots>, a\<^sub>n"} do not have any permanent effect. | |
| 61 | ||
| 62 |   \item @{command "term"}~@{text t} and @{command "prop"}~@{text \<phi>}
 | |
| 63 | read, type-check and print terms or propositions according to the | |
| 64 |   current theory or proof context; the inferred type of @{text t} is
 | |
| 65 | output as well. Note that these commands are also useful in | |
| 66 | inspecting the current environment of term abbreviations. | |
| 67 | ||
| 68 |   \item @{command "typ"}~@{text \<tau>} reads and prints types of the
 | |
| 69 | meta-logic according to the current theory or proof context. | |
| 70 | ||
| 71 |   \item @{command "prf"} displays the (compact) proof term of the
 | |
| 72 | current proof state (if present), or of the given theorems. Note | |
| 73 | that this requires proof terms to be switched on for the current | |
| 74 | object logic (see the ``Proof terms'' section of the Isabelle | |
| 75 | reference manual for information on how to do this). | |
| 76 | ||
| 77 |   \item @{command "full_prf"} is like @{command "prf"}, but displays
 | |
| 78 | the full proof term, i.e.\ also displays information omitted in the | |
| 79 |   compact proof term, which is denoted by ``@{text _}'' placeholders
 | |
| 80 | there. | |
| 81 | ||
| 82 |   \end{description}
 | |
| 83 | ||
| 84 |   All of the diagnostic commands above admit a list of @{text modes}
 | |
| 85 | to be specified, which is appended to the current print mode (see | |
| 86 |   also \cite{isabelle-ref}).  Thus the output behavior may be modified
 | |
| 87 |   according particular print mode features.  For example, @{command
 | |
| 88 |   "pr"}~@{text "(latex xsymbols)"} would print the current proof state
 | |
| 89 | with mathematical symbols and special characters represented in | |
| 90 |   {\LaTeX} source, according to the Isabelle style
 | |
| 91 |   \cite{isabelle-sys}.
 | |
| 92 | ||
| 93 |   Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
 | |
| 94 | systematic way to include formal items into the printed text | |
| 95 | document. | |
| 96 | *} | |
| 97 | ||
| 98 | ||
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 99 | subsection {* Printing limits *}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 100 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 101 | text {*
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 102 |   \begin{mldecls}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 103 |     @{index_ML Pretty.setdepth: "int -> unit"} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 104 |     @{index_ML Pretty.setmargin: "int -> unit"} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 105 |     @{index_ML print_depth: "int -> unit"} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 106 |   \end{mldecls}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 107 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 108 | These ML functions set limits for pretty printed text output. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 109 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 110 |   \begin{description}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 111 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 112 |   \item @{ML Pretty.setdepth}~@{text d} tells the pretty printer to
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 113 |   limit the printing depth to @{text d}.  This affects the display of
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 114 | types, terms, theorems etc. The default value is 0, which permits | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 115 |   printing to an arbitrary depth.  Useful values for @{text d} are 10
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 116 | and 20. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 117 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 118 |   \item @{ML Pretty.setmargin}~@{text m} tells the pretty printer to
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 119 |   assume a right margin (page width) of @{text m}.  The initial margin
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 120 | is 76. User interfaces may adapt this default automatically, when | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 121 | resizing windows etc. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 122 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 123 |   \item @{ML print_depth}~@{text n} limits the printing depth of the
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 124 | ML toplevel pretty printer; the precise effect depends on the ML | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 125 |   compiler and run-time system.  Typically @{text n} should be less
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 126 | than 10. Bigger values such as 100--1000 are useful for debugging. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 127 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 128 |   \end{description}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 129 | *} | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 130 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 131 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 132 | subsection {* Details of printed content *}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 133 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 134 | text {*
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 135 |   \begin{mldecls} 
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 136 |     @{index_ML show_types: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 137 |     @{index_ML show_sorts: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 138 |     @{index_ML show_consts: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 139 |     @{index_ML long_names: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 140 |     @{index_ML short_names: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 141 |     @{index_ML unique_names: "bool ref"} & default @{ML true} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 142 |     @{index_ML show_brackets: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 143 |     @{index_ML eta_contract: "bool ref"} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 144 |     @{index_ML goals_limit: "int ref"} & default @{ML 10} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 145 |     @{index_ML Proof.show_main_goal: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 146 |     @{index_ML show_hyps: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 147 |     @{index_ML show_tags: "bool ref"} & default @{ML false} \\
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 148 |   \end{mldecls}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 149 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 150 | These global ML variables control the detail of information that is | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 151 | displayed for types, terms, theorems, goals etc. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 152 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 153 |   \begin{description}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 154 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 155 |   \item @{ML show_types} and @{ML show_sorts} control printing of type
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 156 | constraints for term variables, and sort constraints for type | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 157 | variables. By default, neither of these are shown in output. If | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 158 |   @{ML show_sorts} is set to @{ML true}, types are always shown as
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 159 | well. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 160 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 161 | Note that displaying types and sorts may explain why a polymorphic | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 162 | inference rule fails to resolve with some goal, or why a rewrite | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 163 | rule does not apply as expected. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 164 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 165 |   \item @{ML show_consts} controls printing of types of constants when
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 166 | printing proof states. Note that the output can be enormous as | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 167 | polymorphic constants often occur at several different type | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 168 | instances. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 169 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 170 |   \item @{ML long_names}, @{ML short_names}, and @{ML unique_names}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 171 | modify the way of printing qualified names in external form. See | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 172 | also the description of document antiquotation options in | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 173 |   \secref{sec:antiq}.
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 174 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 175 |   \item @{ML show_brackets} controls insertion of bracketing in pretty
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 176 |   printed output.  If set to @{ML true}, all sub-expressions of the
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 177 | pretty printing tree will be parenthesized, even if this produces | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 178 | malformed term syntax! This crude way of showing the full structure | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 179 | of pretty printed entities might help to diagnose problems with | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 180 | operator priorities, for example. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 181 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 182 |   \item @{ML eta_contract} controls @{text "\<eta>"}-contracted printing of
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 183 | terms. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 184 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 185 |   The @{text \<eta>}-contraction law asserts @{prop "(\<lambda>x. f x) \<equiv> f"},
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 186 |   provided @{text x} is not free in @{text f}.  It asserts
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 187 |   \emph{extensionality} of functions: @{prop "f \<equiv> g"} if @{prop "f x \<equiv>
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 188 |   g x"} for all @{text x}.  Higher-order unification frequently puts
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 189 |   terms into a fully @{text \<eta>}-expanded form.  For example, if @{text
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 190 |   F} has type @{text "(\<tau> \<Rightarrow> \<tau>) \<Rightarrow> \<tau>"} then its expanded form is @{term
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 191 | "\<lambda>h. F (\<lambda>x. h x)"}. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 192 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 193 |   Setting @{ML eta_contract} makes Isabelle perform @{text
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 194 |   \<eta>}-contractions before printing, so that @{term "\<lambda>h. F (\<lambda>x. h x)"}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 195 |   appears simply as @{text F}.
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 196 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 197 |   Note that the distinction between a term and its @{text \<eta>}-expanded
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 198 | form occasionally matters. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 199 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 200 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 201 |   \item @{ML goals_limit} controls the maximum number of subgoals to
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 202 | be shown in proof state output. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 203 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 204 |   \item @{ML Proof.show_main_goal} controls whether the main result to
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 205 | be proven should be displayed. This information might be relevant | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 206 | for schematic goals, to see the current claim being synthesized. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 207 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 208 |   \item @{ML show_hyps} controls printing of implicit hypotheses of
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 209 | local facts. Normally, only those hypotheses are displayed that are | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 210 |   \emph{not} covered by the assumptions of the current context: this
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 211 | situation indicates a fault in some tool being used. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 212 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 213 |   By setting @{ML show_hyps} to @{ML true}, output of all hypotheses
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 214 | can be enforced. Which is occasionally usefule for diagnostic | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 215 | purposes. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 216 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 217 |   \item @{ML show_tags} controls printing of extra annotations within
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 218 | theorems, such as the case names being attached by the attribute | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 219 |   @{attribute case_names}.
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 220 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 221 |   \end{description}
 | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 222 | *} | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 223 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 224 | |
| 28762 | 225 | section {* Mixfix annotations *}
 | 
| 226 | ||
| 227 | text {* Mixfix annotations specify concrete \emph{inner syntax} of
 | |
| 228 |   Isabelle types and terms.  Some commands such as @{command "types"}
 | |
| 229 |   (see \secref{sec:types-pure}) admit infixes only, while @{command
 | |
| 230 |   "consts"} (see \secref{sec:consts}) and @{command "syntax"} (see
 | |
| 231 |   \secref{sec:syn-trans}) support the full range of general mixfixes
 | |
| 232 | and binders. | |
| 233 | ||
| 234 |   \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
 | |
| 235 |   \begin{rail}
 | |
| 236 |     infix: '(' ('infix' | 'infixl' | 'infixr') string nat ')'
 | |
| 237 | ; | |
| 238 |     mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
 | |
| 239 | ; | |
| 240 |     structmixfix: mixfix | '(' 'structure' ')'
 | |
| 241 | ; | |
| 242 | ||
| 243 | prios: '[' (nat + ',') ']' | |
| 244 | ; | |
| 245 |   \end{rail}
 | |
| 246 | ||
| 247 |   Here the \railtok{string} specifications refer to the actual mixfix
 | |
| 248 | template, which may include literal text, spacing, blocks, and | |
| 249 |   arguments (denoted by ``@{text _}''); the special symbol
 | |
| 250 |   ``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'') represents an index
 | |
| 251 | argument that specifies an implicit structure reference (see also | |
| 252 |   \secref{sec:locale}).  Infix and binder declarations provide common
 | |
| 253 | abbreviations for particular mixfix declarations. So in practice, | |
| 254 | mixfix templates mostly degenerate to literal text for concrete | |
| 255 |   syntax, such as ``@{verbatim "++"}'' for an infix symbol.
 | |
| 256 | ||
| 257 | \medskip In full generality, mixfix declarations work as follows. | |
| 258 |   Suppose a constant @{text "c :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>"} is
 | |
| 259 |   annotated by @{text "(mixfix [p\<^sub>1, \<dots>, p\<^sub>n] p)"}, where @{text
 | |
| 260 |   "mixfix"} is a string @{text "d\<^sub>0 _ d\<^sub>1 _ \<dots> _ d\<^sub>n"} consisting of
 | |
| 261 | delimiters that surround argument positions as indicated by | |
| 262 | underscores. | |
| 263 | ||
| 264 | Altogether this determines a production for a context-free priority | |
| 265 |   grammar, where for each argument @{text "i"} the syntactic category
 | |
| 266 |   is determined by @{text "\<tau>\<^sub>i"} (with priority @{text "p\<^sub>i"}), and
 | |
| 267 |   the result category is determined from @{text "\<tau>"} (with
 | |
| 268 |   priority @{text "p"}).  Priority specifications are optional, with
 | |
| 269 | default 0 for arguments and 1000 for the result. | |
| 270 | ||
| 271 |   Since @{text "\<tau>"} may be again a function type, the constant
 | |
| 272 | type scheme may have more argument positions than the mixfix | |
| 273 |   pattern.  Printing a nested application @{text "c t\<^sub>1 \<dots> t\<^sub>m"} for
 | |
| 274 |   @{text "m > n"} works by attaching concrete notation only to the
 | |
| 275 |   innermost part, essentially by printing @{text "(c t\<^sub>1 \<dots> t\<^sub>n) \<dots> t\<^sub>m"}
 | |
| 276 | instead. If a term has fewer arguments than specified in the mixfix | |
| 277 | template, the concrete syntax is ignored. | |
| 278 | ||
| 279 | \medskip A mixfix template may also contain additional directives | |
| 280 | for pretty printing, notably spaces, blocks, and breaks. The | |
| 281 | general template format is a sequence over any of the following | |
| 282 | entities. | |
| 283 | ||
| 284 |   \begin{itemize}
 | |
| 285 | ||
| 286 |   \item @{text "\<^bold>d"} is a delimiter, namely a non-empty
 | |
| 287 |   sequence of characters other than the special characters @{text "'"}
 | |
| 288 |   (single quote), @{text "_"} (underscore), @{text "\<index>"} (index
 | |
| 289 |   symbol), @{text "/"} (slash), @{text "("} and @{text ")"}
 | |
| 290 | (parentheses). | |
| 291 | ||
| 292 | A single quote escapes the special meaning of these meta-characters, | |
| 293 | producing a literal version of the following character, unless that | |
| 294 | is a blank. A single quote followed by a blank separates | |
| 295 | delimiters, without affecting printing, but input tokens may have | |
| 296 | additional white space here. | |
| 297 | ||
| 298 |   \item @{text "_"} is an argument position, which stands for a
 | |
| 299 | certain syntactic category in the underlying grammar. | |
| 300 | ||
| 301 |   \item @{text "\<index>"} is an indexed argument position; this is
 | |
| 302 | the place where implicit structure arguments can be attached. | |
| 303 | ||
| 304 |   \item @{text "\<^bold>s"} is a non-empty sequence of spaces for
 | |
| 305 | printing. This and the following specifications do not affect | |
| 306 | parsing at all. | |
| 307 | ||
| 308 |   \item @{text "(\<^bold>n"} opens a pretty printing block.  The
 | |
| 309 | optional number specifies how much indentation to add when a line | |
| 310 | break occurs within the block. If the parenthesis is not followed | |
| 311 | by digits, the indentation defaults to 0. A block specified via | |
| 312 |   @{text "(00"} is unbreakable.
 | |
| 313 | ||
| 314 |   \item @{text ")"} closes a pretty printing block.
 | |
| 315 | ||
| 316 |   \item @{text "//"} forces a line break.
 | |
| 317 | ||
| 318 |   \item @{text "/\<^bold>s"} allows a line break.  Here @{text
 | |
| 319 | "\<^bold>s"} stands for the string of spaces (zero or more) right | |
| 320 | after the slash. These spaces are printed if the break is | |
| 321 |   \emph{not} taken.
 | |
| 322 | ||
| 323 |   \end{itemize}
 | |
| 324 | ||
| 325 |   For example, the template @{text "(_ +/ _)"} specifies an infix
 | |
| 326 |   operator.  There are two argument positions; the delimiter @{text
 | |
| 327 | "+"} is preceded by a space and followed by a space or line break; | |
| 328 | the entire phrase is a pretty printing block. | |
| 329 | ||
| 330 | The general idea of pretty printing with blocks and breaks is also | |
| 331 |   described in \cite{paulson-ml2}.
 | |
| 332 | *} | |
| 333 | ||
| 334 | ||
| 335 | section {* Additional term notation *}
 | |
| 336 | ||
| 337 | text {*
 | |
| 338 |   \begin{matharray}{rcll}
 | |
| 339 |     @{command_def "notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | |
| 340 |     @{command_def "no_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | |
| 341 |   \end{matharray}
 | |
| 342 | ||
| 343 |   \begin{rail}
 | |
| 344 |     ('notation' | 'no\_notation') target? mode? (nameref structmixfix + 'and')
 | |
| 345 | ; | |
| 346 |   \end{rail}
 | |
| 347 | ||
| 348 |   \begin{description}
 | |
| 349 | ||
| 350 |   \item @{command "notation"}~@{text "c (mx)"} associates mixfix
 | |
| 351 | syntax with an existing constant or fixed variable. This is a | |
| 352 |   robust interface to the underlying @{command "syntax"} primitive
 | |
| 353 |   (\secref{sec:syn-trans}).  Type declaration and internal syntactic
 | |
| 354 | representation of the given entity is retrieved from the context. | |
| 355 | ||
| 356 |   \item @{command "no_notation"} is similar to @{command "notation"},
 | |
| 357 | but removes the specified syntax annotation from the present | |
| 358 | context. | |
| 359 | ||
| 360 |   \end{description}
 | |
| 361 | *} | |
| 362 | ||
| 363 | section {* Syntax and translations \label{sec:syn-trans} *}
 | |
| 364 | ||
| 365 | text {*
 | |
| 366 |   \begin{matharray}{rcl}
 | |
| 367 |     @{command_def "nonterminals"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 368 |     @{command_def "syntax"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 369 |     @{command_def "no_syntax"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 370 |     @{command_def "translations"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 371 |     @{command_def "no_translations"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 372 |   \end{matharray}
 | |
| 373 | ||
| 374 |   \begin{rail}
 | |
| 375 | 'nonterminals' (name +) | |
| 376 | ; | |
| 377 |     ('syntax' | 'no\_syntax') mode? (constdecl +)
 | |
| 378 | ; | |
| 379 |     ('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
 | |
| 380 | ; | |
| 381 | ||
| 382 |     mode: ('(' ( name | 'output' | name 'output' ) ')')
 | |
| 383 | ; | |
| 384 |     transpat: ('(' nameref ')')? string
 | |
| 385 | ; | |
| 386 |   \end{rail}
 | |
| 387 | ||
| 388 |   \begin{description}
 | |
| 389 | ||
| 390 |   \item @{command "nonterminals"}~@{text c} declares a type
 | |
| 391 |   constructor @{text c} (without arguments) to act as purely syntactic
 | |
| 392 | type: a nonterminal symbol of the inner syntax. | |
| 393 | ||
| 394 |   \item @{command "syntax"}~@{text "(mode) decls"} is similar to
 | |
| 395 |   @{command "consts"}~@{text decls}, except that the actual logical
 | |
| 396 | signature extension is omitted. Thus the context free grammar of | |
| 397 | Isabelle's inner syntax may be augmented in arbitrary ways, | |
| 398 |   independently of the logic.  The @{text mode} argument refers to the
 | |
| 399 |   print mode that the grammar rules belong; unless the @{keyword_ref
 | |
| 400 | "output"} indicator is given, all productions are added both to the | |
| 401 | input and output grammar. | |
| 402 | ||
| 403 |   \item @{command "no_syntax"}~@{text "(mode) decls"} removes grammar
 | |
| 404 |   declarations (and translations) resulting from @{text decls}, which
 | |
| 405 |   are interpreted in the same manner as for @{command "syntax"} above.
 | |
| 406 | ||
| 407 |   \item @{command "translations"}~@{text rules} specifies syntactic
 | |
| 408 |   translation rules (i.e.\ macros): parse~/ print rules (@{text "\<rightleftharpoons>"}),
 | |
| 409 |   parse rules (@{text "\<rightharpoonup>"}), or print rules (@{text "\<leftharpoondown>"}).
 | |
| 410 | Translation patterns may be prefixed by the syntactic category to be | |
| 411 |   used for parsing; the default is @{text logic}.
 | |
| 412 | ||
| 413 |   \item @{command "no_translations"}~@{text rules} removes syntactic
 | |
| 414 | translation rules, which are interpreted in the same manner as for | |
| 415 |   @{command "translations"} above.
 | |
| 416 | ||
| 417 |   \end{description}
 | |
| 418 | *} | |
| 419 | ||
| 420 | ||
| 421 | section {* Syntax translation functions *}
 | |
| 422 | ||
| 423 | text {*
 | |
| 424 |   \begin{matharray}{rcl}
 | |
| 425 |     @{command_def "parse_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 426 |     @{command_def "parse_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 427 |     @{command_def "print_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 428 |     @{command_def "typed_print_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 429 |     @{command_def "print_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | |
| 430 |   \end{matharray}
 | |
| 431 | ||
| 432 |   \begin{rail}
 | |
| 433 | ( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' | | |
| 434 |     'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text
 | |
| 435 | ; | |
| 436 |   \end{rail}
 | |
| 437 | ||
| 438 | Syntax translation functions written in ML admit almost arbitrary | |
| 439 | manipulations of Isabelle's inner syntax. Any of the above commands | |
| 440 |   have a single \railqtok{text} argument that refers to an ML
 | |
| 441 | expression of appropriate type, which are as follows by default: | |
| 442 | ||
| 443 | %FIXME proper antiquotations | |
| 444 | \begin{ttbox}
 | |
| 445 | val parse_ast_translation : (string * (ast list -> ast)) list | |
| 446 | val parse_translation : (string * (term list -> term)) list | |
| 447 | val print_translation : (string * (term list -> term)) list | |
| 448 | val typed_print_translation : | |
| 449 | (string * (bool -> typ -> term list -> term)) list | |
| 450 | val print_ast_translation : (string * (ast list -> ast)) list | |
| 451 | \end{ttbox}
 | |
| 452 | ||
| 453 |   If the @{text "(advanced)"} option is given, the corresponding
 | |
| 454 | translation functions may depend on the current theory or proof | |
| 455 | context. This allows to implement advanced syntax mechanisms, as | |
| 456 | translations functions may refer to specific theory declarations or | |
| 457 | auxiliary proof data. | |
| 458 | ||
| 459 |   See also \cite[\S8]{isabelle-ref} for more information on the
 | |
| 460 | general concept of syntax transformations in Isabelle. | |
| 461 | ||
| 462 | %FIXME proper antiquotations | |
| 463 | \begin{ttbox}
 | |
| 464 | val parse_ast_translation: | |
| 465 | (string * (Proof.context -> ast list -> ast)) list | |
| 466 | val parse_translation: | |
| 467 | (string * (Proof.context -> term list -> term)) list | |
| 468 | val print_translation: | |
| 469 | (string * (Proof.context -> term list -> term)) list | |
| 470 | val typed_print_translation: | |
| 471 | (string * (Proof.context -> bool -> typ -> term list -> term)) list | |
| 472 | val print_ast_translation: | |
| 473 | (string * (Proof.context -> ast list -> ast)) list | |
| 474 | \end{ttbox}
 | |
| 475 | *} | |
| 476 | ||
| 477 | end |