| author | bulwahn | 
| Sun, 03 Jul 2011 09:59:25 +0200 | |
| changeset 43656 | 9ece73262746 | 
| parent 42926 | a8b655d089ac | 
| child 45703 | c7a13ce60161 | 
| permissions | -rw-r--r-- | 
| 28762 | 1  | 
theory Inner_Syntax  | 
| 42651 | 2  | 
imports Base Main  | 
| 28762 | 3  | 
begin  | 
4  | 
||
| 28778 | 5  | 
chapter {* Inner syntax --- the term language \label{ch:inner-syntax} *}
 | 
| 28762 | 6  | 
|
7  | 
section {* Printing logical entities *}
 | 
|
8  | 
||
9  | 
subsection {* Diagnostic commands *}
 | 
|
10  | 
||
11  | 
text {*
 | 
|
12  | 
  \begin{matharray}{rcl}
 | 
|
| 
28766
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
13  | 
    @{command_def "typ"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 28762 | 14  | 
    @{command_def "term"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
15  | 
    @{command_def "prop"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
|
| 
28766
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
16  | 
    @{command_def "thm"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 28762 | 17  | 
    @{command_def "prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
18  | 
    @{command_def "full_prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
|
| 
28766
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
19  | 
    @{command_def "pr"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
 | 
| 28762 | 20  | 
  \end{matharray}
 | 
21  | 
||
22  | 
These diagnostic commands assist interactive development by printing  | 
|
23  | 
internal logical entities in a human-readable fashion.  | 
|
24  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
25  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
26  | 
    @@{command typ} @{syntax modes}? @{syntax type}
 | 
| 28762 | 27  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
28  | 
    @@{command term} @{syntax modes}? @{syntax term}
 | 
| 28762 | 29  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
30  | 
    @@{command prop} @{syntax modes}? @{syntax prop}
 | 
| 28762 | 31  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
32  | 
    @@{command thm} @{syntax modes}? @{syntax thmrefs}
 | 
| 28762 | 33  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
34  | 
    ( @@{command prf} | @@{command full_prf} ) @{syntax modes}? @{syntax thmrefs}?
 | 
| 28762 | 35  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
36  | 
    @@{command pr} @{syntax modes}? @{syntax nat}?
 | 
| 28762 | 37  | 
;  | 
38  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
39  | 
    @{syntax_def modes}: '(' (@{syntax name} + ) ')'
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
40  | 
"}  | 
| 28762 | 41  | 
|
42  | 
  \begin{description}
 | 
|
43  | 
||
| 
28766
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
44  | 
  \item @{command "typ"}~@{text \<tau>} reads and prints types of the
 | 
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
45  | 
meta-logic according to the current theory or proof context.  | 
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
46  | 
|
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
47  | 
  \item @{command "term"}~@{text t} and @{command "prop"}~@{text \<phi>}
 | 
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
48  | 
read, type-check and print terms or propositions according to the  | 
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
49  | 
  current theory or proof context; the inferred type of @{text t} is
 | 
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
50  | 
output as well. Note that these commands are also useful in  | 
| 
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
51  | 
inspecting the current environment of term abbreviations.  | 
| 28762 | 52  | 
|
53  | 
  \item @{command "thm"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} retrieves
 | 
|
54  | 
theorems from the current theory or proof context. Note that any  | 
|
55  | 
attributes included in the theorem specifications are applied to a  | 
|
56  | 
temporary context derived from the current theory or proof; the  | 
|
57  | 
  result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1,
 | 
|
58  | 
\<dots>, a\<^sub>n"} do not have any permanent effect.  | 
|
59  | 
||
60  | 
  \item @{command "prf"} displays the (compact) proof term of the
 | 
|
61  | 
current proof state (if present), or of the given theorems. Note  | 
|
62  | 
that this requires proof terms to be switched on for the current  | 
|
63  | 
object logic (see the ``Proof terms'' section of the Isabelle  | 
|
64  | 
reference manual for information on how to do this).  | 
|
65  | 
||
66  | 
  \item @{command "full_prf"} is like @{command "prf"}, but displays
 | 
|
67  | 
the full proof term, i.e.\ also displays information omitted in the  | 
|
68  | 
  compact proof term, which is denoted by ``@{text _}'' placeholders
 | 
|
69  | 
there.  | 
|
70  | 
||
| 39165 | 71  | 
  \item @{command "pr"}~@{text "goals"} prints the current proof state
 | 
72  | 
(if present), including current facts and goals. The optional limit  | 
|
73  | 
arguments affect the number of goals to be displayed, which is  | 
|
74  | 
initially 10. Omitting limit value leaves the current setting  | 
|
75  | 
unchanged.  | 
|
| 
28766
 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28765 
diff
changeset
 | 
76  | 
|
| 28762 | 77  | 
  \end{description}
 | 
78  | 
||
79  | 
  All of the diagnostic commands above admit a list of @{text modes}
 | 
|
| 42926 | 80  | 
to be specified, which is appended to the current print mode; see  | 
81  | 
  \secref{sec:print-modes}.  Thus the output behavior may be modified
 | 
|
| 28762 | 82  | 
  according particular print mode features.  For example, @{command
 | 
83  | 
  "pr"}~@{text "(latex xsymbols)"} would print the current proof state
 | 
|
84  | 
with mathematical symbols and special characters represented in  | 
|
85  | 
  {\LaTeX} source, according to the Isabelle style
 | 
|
86  | 
  \cite{isabelle-sys}.
 | 
|
87  | 
||
88  | 
  Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
 | 
|
89  | 
systematic way to include formal items into the printed text  | 
|
90  | 
document.  | 
|
91  | 
*}  | 
|
92  | 
||
93  | 
||
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
94  | 
subsection {* Details of printed content *}
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
95  | 
|
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
96  | 
text {*
 | 
| 42655 | 97  | 
  \begin{tabular}{rcll}
 | 
98  | 
    @{attribute_def show_types} & : & @{text attribute} & default @{text false} \\
 | 
|
99  | 
    @{attribute_def show_sorts} & : & @{text attribute} & default @{text false} \\
 | 
|
100  | 
    @{attribute_def show_consts} & : & @{text attribute} & default @{text false} \\
 | 
|
101  | 
    @{attribute_def show_abbrevs} & : & @{text attribute} & default @{text true} \\
 | 
|
102  | 
    @{attribute_def show_brackets} & : & @{text attribute} & default @{text false} \\
 | 
|
| 
42669
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42655 
diff
changeset
 | 
103  | 
    @{attribute_def names_long} & : & @{text attribute} & default @{text false} \\
 | 
| 
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42655 
diff
changeset
 | 
104  | 
    @{attribute_def names_short} & : & @{text attribute} & default @{text false} \\
 | 
| 
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42655 
diff
changeset
 | 
105  | 
    @{attribute_def names_unique} & : & @{text attribute} & default @{text true} \\
 | 
| 42655 | 106  | 
    @{attribute_def eta_contract} & : & @{text attribute} & default @{text true} \\
 | 
107  | 
    @{attribute_def goals_limit} & : & @{text attribute} & default @{text 10} \\
 | 
|
108  | 
    @{attribute_def show_main_goal} & : & @{text attribute} & default @{text false} \\
 | 
|
109  | 
    @{attribute_def show_hyps} & : & @{text attribute} & default @{text false} \\
 | 
|
110  | 
    @{attribute_def show_tags} & : & @{text attribute} & default @{text false} \\
 | 
|
111  | 
    @{attribute_def show_question_marks} & : & @{text attribute} & default @{text true} \\
 | 
|
112  | 
  \end{tabular}
 | 
|
113  | 
\medskip  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
114  | 
|
| 42655 | 115  | 
These configuration options control the detail of information that  | 
116  | 
is displayed for types, terms, theorems, goals etc. See also  | 
|
117  | 
  \secref{sec:config}.
 | 
|
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
118  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
119  | 
  \begin{description}
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
120  | 
|
| 42655 | 121  | 
  \item @{attribute show_types} and @{attribute show_sorts} control
 | 
122  | 
printing of type constraints for term variables, and sort  | 
|
123  | 
constraints for type variables. By default, neither of these are  | 
|
124  | 
  shown in output.  If @{attribute show_sorts} is enabled, types are
 | 
|
125  | 
always shown as well.  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
126  | 
|
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
127  | 
Note that displaying types and sorts may explain why a polymorphic  | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
128  | 
inference rule fails to resolve with some goal, or why a rewrite  | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
129  | 
rule does not apply as expected.  | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
130  | 
|
| 42655 | 131  | 
  \item @{attribute show_consts} controls printing of types of
 | 
132  | 
constants when displaying a goal state.  | 
|
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
133  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
134  | 
Note that the output can be enormous, because polymorphic constants  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
135  | 
often occur at several different type instances.  | 
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
136  | 
|
| 42655 | 137  | 
  \item @{attribute show_abbrevs} controls folding of constant
 | 
138  | 
abbreviations.  | 
|
| 
40879
 
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
 
wenzelm 
parents: 
40255 
diff
changeset
 | 
139  | 
|
| 42655 | 140  | 
  \item @{attribute show_brackets} controls bracketing in pretty
 | 
141  | 
printed output. If enabled, all sub-expressions of the pretty  | 
|
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
142  | 
printing tree will be parenthesized, even if this produces malformed  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
143  | 
term syntax! This crude way of showing the internal structure of  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
144  | 
pretty printed entities may occasionally help to diagnose problems  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
145  | 
with operator priorities, for example.  | 
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
146  | 
|
| 
42669
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42655 
diff
changeset
 | 
147  | 
  \item @{attribute names_long}, @{attribute names_short}, and
 | 
| 
 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 
wenzelm 
parents: 
42655 
diff
changeset
 | 
148  | 
  @{attribute names_unique} control the way of printing fully
 | 
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
42279 
diff
changeset
 | 
149  | 
qualified internal names in external form. See also  | 
| 
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
42279 
diff
changeset
 | 
150  | 
  \secref{sec:antiq} for the document antiquotation options of the
 | 
| 
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
42279 
diff
changeset
 | 
151  | 
same names.  | 
| 
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
42279 
diff
changeset
 | 
152  | 
|
| 42655 | 153  | 
  \item @{attribute eta_contract} controls @{text "\<eta>"}-contracted
 | 
154  | 
printing of terms.  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
155  | 
|
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
156  | 
  The @{text \<eta>}-contraction law asserts @{prop "(\<lambda>x. f x) \<equiv> f"},
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
157  | 
  provided @{text x} is not free in @{text f}.  It asserts
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
158  | 
  \emph{extensionality} of functions: @{prop "f \<equiv> g"} if @{prop "f x \<equiv>
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
159  | 
  g x"} for all @{text x}.  Higher-order unification frequently puts
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
160  | 
  terms into a fully @{text \<eta>}-expanded form.  For example, if @{text
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
161  | 
  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: 
28762 
diff
changeset
 | 
162  | 
"\<lambda>h. F (\<lambda>x. h x)"}.  | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
163  | 
|
| 42655 | 164  | 
  Enabling @{attribute eta_contract} makes Isabelle perform @{text
 | 
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
165  | 
  \<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: 
28762 
diff
changeset
 | 
166  | 
  appears simply as @{text F}.
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
167  | 
|
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
168  | 
  Note that the distinction between a term and its @{text \<eta>}-expanded
 | 
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
169  | 
form occasionally matters. While higher-order resolution and  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
170  | 
  rewriting operate modulo @{text "\<alpha>\<beta>\<eta>"}-conversion, some other tools
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
171  | 
might look at terms more discretely.  | 
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
172  | 
|
| 42655 | 173  | 
  \item @{attribute goals_limit} controls the maximum number of
 | 
| 39130 | 174  | 
subgoals to be shown in goal output.  | 
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
175  | 
|
| 42655 | 176  | 
  \item @{attribute show_main_goal} controls whether the main result
 | 
177  | 
to be proven should be displayed. This information might be  | 
|
| 39130 | 178  | 
relevant for schematic goals, to inspect the current claim that has  | 
179  | 
been synthesized so far.  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
180  | 
|
| 42655 | 181  | 
  \item @{attribute show_hyps} controls printing of implicit
 | 
182  | 
hypotheses of local facts. Normally, only those hypotheses are  | 
|
183  | 
  displayed that are \emph{not} covered by the assumptions of the
 | 
|
184  | 
current context: this situation indicates a fault in some tool being  | 
|
185  | 
used.  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
186  | 
|
| 42655 | 187  | 
  By enabling @{attribute show_hyps}, output of \emph{all} hypotheses
 | 
188  | 
can be enforced, which is occasionally useful for diagnostic  | 
|
189  | 
purposes.  | 
|
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
190  | 
|
| 42655 | 191  | 
  \item @{attribute show_tags} controls printing of extra annotations
 | 
192  | 
within theorems, such as internal position information, or the case  | 
|
193  | 
  names being attached by the attribute @{attribute case_names}.
 | 
|
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
194  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
195  | 
  Note that the @{attribute tagged} and @{attribute untagged}
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
196  | 
attributes provide low-level access to the collection of tags  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
197  | 
associated with a theorem.  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
198  | 
|
| 42655 | 199  | 
  \item @{attribute show_question_marks} controls printing of question
 | 
200  | 
  marks for schematic variables, such as @{text ?x}.  Only the leading
 | 
|
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
201  | 
question mark is affected, the remaining text is unchanged  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
202  | 
(including proper markup for schematic variables that might be  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
203  | 
relevant for user interfaces).  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
204  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
205  | 
  \end{description}
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
206  | 
*}  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
207  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
208  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
209  | 
subsection {* Printing limits *}
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
210  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
211  | 
text {*
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
212  | 
  \begin{mldecls}
 | 
| 36745 | 213  | 
    @{index_ML Pretty.margin_default: "int Unsynchronized.ref"} \\
 | 
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
214  | 
    @{index_ML print_depth: "int -> unit"} \\
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
215  | 
  \end{mldecls}
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
216  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
217  | 
These ML functions set limits for pretty printed text.  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
218  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
219  | 
  \begin{description}
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
220  | 
|
| 36745 | 221  | 
  \item @{ML Pretty.margin_default} indicates the global default for
 | 
222  | 
the right margin of the built-in pretty printer, with initial value  | 
|
223  | 
76. Note that user-interfaces typically control margins  | 
|
224  | 
automatically when resizing windows, or even bypass the formatting  | 
|
225  | 
engine of Isabelle/ML altogether and do it within the front end via  | 
|
226  | 
Isabelle/Scala.  | 
|
| 
28765
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
227  | 
|
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
228  | 
  \item @{ML print_depth}~@{text n} limits the printing depth of the
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
229  | 
ML toplevel pretty printer; the precise effect depends on the ML  | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
230  | 
  compiler and run-time system.  Typically @{text n} should be less
 | 
| 
 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 
wenzelm 
parents: 
28763 
diff
changeset
 | 
231  | 
than 10. Bigger values such as 100--1000 are useful for debugging.  | 
| 
28763
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
232  | 
|
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
233  | 
  \end{description}
 | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
234  | 
*}  | 
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
235  | 
|
| 
 
b5e6122ff575
added pretty printing options (from old ref manual);
 
wenzelm 
parents: 
28762 
diff
changeset
 | 
236  | 
|
| 28762 | 237  | 
section {* Mixfix annotations *}
 | 
238  | 
||
239  | 
text {* Mixfix annotations specify concrete \emph{inner syntax} of
 | 
|
| 
35351
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
32833 
diff
changeset
 | 
240  | 
Isabelle types and terms. Locally fixed parameters in toplevel  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
32833 
diff
changeset
 | 
241  | 
theorem statements, locale specifications etc.\ also admit mixfix  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
32833 
diff
changeset
 | 
242  | 
annotations.  | 
| 28762 | 243  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
244  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
245  | 
    @{syntax_def \"infix\"}: '(' (@'infix' | @'infixl' | @'infixr')
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
246  | 
      @{syntax string} @{syntax nat} ')'
 | 
| 28762 | 247  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
248  | 
    @{syntax_def mixfix}: @{syntax \"infix\"} | '(' @{syntax string} prios? @{syntax nat}? ')' |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
249  | 
    '(' @'binder' @{syntax string} prios? @{syntax nat} ')'
 | 
| 28762 | 250  | 
;  | 
| 42705 | 251  | 
    @{syntax_def struct_mixfix}: @{syntax mixfix} | '(' @'structure' ')'
 | 
| 28762 | 252  | 
;  | 
253  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
254  | 
    prios: '[' (@{syntax nat} + ',') ']'
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
255  | 
"}  | 
| 28762 | 256  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
257  | 
  Here the @{syntax string} specifications refer to the actual mixfix
 | 
| 28762 | 258  | 
template, which may include literal text, spacing, blocks, and  | 
259  | 
  arguments (denoted by ``@{text _}''); the special symbol
 | 
|
260  | 
  ``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'') represents an index
 | 
|
261  | 
argument that specifies an implicit structure reference (see also  | 
|
262  | 
  \secref{sec:locale}).  Infix and binder declarations provide common
 | 
|
263  | 
abbreviations for particular mixfix declarations. So in practice,  | 
|
264  | 
mixfix templates mostly degenerate to literal text for concrete  | 
|
265  | 
  syntax, such as ``@{verbatim "++"}'' for an infix symbol.
 | 
|
266  | 
||
267  | 
\medskip In full generality, mixfix declarations work as follows.  | 
|
268  | 
  Suppose a constant @{text "c :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>"} is
 | 
|
269  | 
  annotated by @{text "(mixfix [p\<^sub>1, \<dots>, p\<^sub>n] p)"}, where @{text
 | 
|
270  | 
  "mixfix"} is a string @{text "d\<^sub>0 _ d\<^sub>1 _ \<dots> _ d\<^sub>n"} consisting of
 | 
|
271  | 
delimiters that surround argument positions as indicated by  | 
|
272  | 
underscores.  | 
|
273  | 
||
274  | 
Altogether this determines a production for a context-free priority  | 
|
275  | 
  grammar, where for each argument @{text "i"} the syntactic category
 | 
|
276  | 
  is determined by @{text "\<tau>\<^sub>i"} (with priority @{text "p\<^sub>i"}), and
 | 
|
277  | 
  the result category is determined from @{text "\<tau>"} (with
 | 
|
278  | 
  priority @{text "p"}).  Priority specifications are optional, with
 | 
|
279  | 
default 0 for arguments and 1000 for the result.  | 
|
280  | 
||
281  | 
  Since @{text "\<tau>"} may be again a function type, the constant
 | 
|
282  | 
type scheme may have more argument positions than the mixfix  | 
|
283  | 
  pattern.  Printing a nested application @{text "c t\<^sub>1 \<dots> t\<^sub>m"} for
 | 
|
284  | 
  @{text "m > n"} works by attaching concrete notation only to the
 | 
|
285  | 
  innermost part, essentially by printing @{text "(c t\<^sub>1 \<dots> t\<^sub>n) \<dots> t\<^sub>m"}
 | 
|
286  | 
instead. If a term has fewer arguments than specified in the mixfix  | 
|
287  | 
template, the concrete syntax is ignored.  | 
|
288  | 
||
289  | 
\medskip A mixfix template may also contain additional directives  | 
|
290  | 
for pretty printing, notably spaces, blocks, and breaks. The  | 
|
291  | 
general template format is a sequence over any of the following  | 
|
292  | 
entities.  | 
|
293  | 
||
| 28778 | 294  | 
  \begin{description}
 | 
| 28762 | 295  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
296  | 
  \item @{text "d"} is a delimiter, namely a non-empty sequence of
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
297  | 
characters other than the following special characters:  | 
| 28762 | 298  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
299  | 
\smallskip  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
300  | 
  \begin{tabular}{ll}
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
301  | 
    @{verbatim "'"} & single quote \\
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
302  | 
    @{verbatim "_"} & underscore \\
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
303  | 
    @{text "\<index>"} & index symbol \\
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
304  | 
    @{verbatim "("} & open parenthesis \\
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
305  | 
    @{verbatim ")"} & close parenthesis \\
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
306  | 
    @{verbatim "/"} & slash \\
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
307  | 
  \end{tabular}
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
308  | 
\medskip  | 
| 28762 | 309  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
310  | 
  \item @{verbatim "'"} escapes the special meaning of these
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
311  | 
meta-characters, producing a literal version of the following  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
312  | 
character, unless that is a blank.  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
313  | 
|
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
314  | 
A single quote followed by a blank separates delimiters, without  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
315  | 
affecting printing, but input tokens may have additional white space  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
316  | 
here.  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
317  | 
|
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
318  | 
  \item @{verbatim "_"} is an argument position, which stands for a
 | 
| 28762 | 319  | 
certain syntactic category in the underlying grammar.  | 
320  | 
||
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
321  | 
  \item @{text "\<index>"} is an indexed argument position; this is the place
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
322  | 
where implicit structure arguments can be attached.  | 
| 28762 | 323  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
324  | 
  \item @{text "s"} is a non-empty sequence of spaces for printing.
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
325  | 
This and the following specifications do not affect parsing at all.  | 
| 28762 | 326  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
327  | 
  \item @{verbatim "("}@{text n} opens a pretty printing block.  The
 | 
| 28762 | 328  | 
optional number specifies how much indentation to add when a line  | 
329  | 
break occurs within the block. If the parenthesis is not followed  | 
|
330  | 
by digits, the indentation defaults to 0. A block specified via  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
331  | 
  @{verbatim "(00"} is unbreakable.
 | 
| 28762 | 332  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
333  | 
  \item @{verbatim ")"} closes a pretty printing block.
 | 
| 28762 | 334  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
335  | 
  \item @{verbatim "//"} forces a line break.
 | 
| 28762 | 336  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
337  | 
  \item @{verbatim "/"}@{text s} allows a line break.  Here @{text s}
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
338  | 
stands for the string of spaces (zero or more) right after the  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
339  | 
  slash.  These spaces are printed if the break is \emph{not} taken.
 | 
| 28762 | 340  | 
|
| 28778 | 341  | 
  \end{description}
 | 
| 28762 | 342  | 
|
| 
28771
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
343  | 
  For example, the template @{verbatim "(_ +/ _)"} specifies an infix
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
344  | 
operator. There are two argument positions; the delimiter  | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
345  | 
  @{verbatim "+"} is preceded by a space and followed by a space or
 | 
| 
 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 
wenzelm 
parents: 
28770 
diff
changeset
 | 
346  | 
line break; the entire phrase is a pretty printing block.  | 
| 28762 | 347  | 
|
348  | 
The general idea of pretty printing with blocks and breaks is also  | 
|
349  | 
  described in \cite{paulson-ml2}.
 | 
|
350  | 
*}  | 
|
351  | 
||
352  | 
||
| 35413 | 353  | 
section {* Explicit notation *}
 | 
| 28762 | 354  | 
|
355  | 
text {*
 | 
|
356  | 
  \begin{matharray}{rcll}
 | 
|
| 35413 | 357  | 
    @{command_def "type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
358  | 
    @{command_def "no_type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
| 28762 | 359  | 
    @{command_def "notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
360  | 
    @{command_def "no_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
|
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
361  | 
    @{command_def "write"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
 | 
| 28762 | 362  | 
  \end{matharray}
 | 
363  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
364  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
365  | 
    (@@{command type_notation} | @@{command no_type_notation}) @{syntax target}?
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
366  | 
      @{syntax mode}? \\ (@{syntax nameref} @{syntax mixfix} + @'and')
 | 
| 35413 | 367  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
368  | 
    (@@{command notation} | @@{command no_notation}) @{syntax target}? @{syntax mode}? \\
 | 
| 42705 | 369  | 
      (@{syntax nameref} @{syntax struct_mixfix} + @'and')
 | 
| 28762 | 370  | 
;  | 
| 42705 | 371  | 
    @@{command write} @{syntax mode}? (@{syntax nameref} @{syntax struct_mixfix} + @'and')
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
372  | 
"}  | 
| 28762 | 373  | 
|
374  | 
  \begin{description}
 | 
|
375  | 
||
| 35413 | 376  | 
  \item @{command "type_notation"}~@{text "c (mx)"} associates mixfix
 | 
377  | 
syntax with an existing type constructor. The arity of the  | 
|
378  | 
constructor is retrieved from the context.  | 
|
379  | 
||
380  | 
  \item @{command "no_type_notation"} is similar to @{command
 | 
|
381  | 
"type_notation"}, but removes the specified syntax annotation from  | 
|
382  | 
the present context.  | 
|
383  | 
||
| 28762 | 384  | 
  \item @{command "notation"}~@{text "c (mx)"} associates mixfix
 | 
| 35413 | 385  | 
syntax with an existing constant or fixed variable. The type  | 
386  | 
declaration of the given entity is retrieved from the context.  | 
|
| 28762 | 387  | 
|
388  | 
  \item @{command "no_notation"} is similar to @{command "notation"},
 | 
|
389  | 
but removes the specified syntax annotation from the present  | 
|
390  | 
context.  | 
|
391  | 
||
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
392  | 
  \item @{command "write"} is similar to @{command "notation"}, but
 | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
393  | 
works within an Isar proof body.  | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
394  | 
|
| 28762 | 395  | 
  \end{description}
 | 
| 35413 | 396  | 
|
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
397  | 
  Note that the more primitive commands @{command "syntax"} and
 | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
398  | 
  @{command "no_syntax"} (\secref{sec:syn-trans}) provide raw access
 | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35413 
diff
changeset
 | 
399  | 
to the syntax tables of a global theory.  | 
| 28762 | 400  | 
*}  | 
401  | 
||
| 28778 | 402  | 
|
403  | 
section {* The Pure syntax \label{sec:pure-syntax} *}
 | 
|
| 
28769
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
404  | 
|
| 28777 | 405  | 
subsection {* Priority grammars \label{sec:priority-grammar} *}
 | 
| 
28769
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
406  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
407  | 
text {* A context-free grammar consists of a set of \emph{terminal
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
408  | 
  symbols}, a set of \emph{nonterminal symbols} and a set of
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
409  | 
  \emph{productions}.  Productions have the form @{text "A = \<gamma>"},
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
410  | 
  where @{text A} is a nonterminal and @{text \<gamma>} is a string of
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
411  | 
terminals and nonterminals. One designated nonterminal is called  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
412  | 
  the \emph{root symbol}.  The language defined by the grammar
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
413  | 
consists of all strings of terminals that can be derived from the  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
414  | 
root symbol by applying productions as rewrite rules.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
415  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
416  | 
  The standard Isabelle parser for inner syntax uses a \emph{priority
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
417  | 
grammar}. Each nonterminal is decorated by an integer priority:  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
418  | 
  @{text "A\<^sup>(\<^sup>p\<^sup>)"}.  In a derivation, @{text "A\<^sup>(\<^sup>p\<^sup>)"} may be rewritten
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
419  | 
  using a production @{text "A\<^sup>(\<^sup>q\<^sup>) = \<gamma>"} only if @{text "p \<le> q"}.  Any
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
420  | 
priority grammar can be translated into a normal context-free  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
421  | 
grammar by introducing new nonterminals and productions.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
422  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
423  | 
  \medskip Formally, a set of context free productions @{text G}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
424  | 
  induces a derivation relation @{text "\<longrightarrow>\<^sub>G"} as follows.  Let @{text
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
425  | 
  \<alpha>} and @{text \<beta>} denote strings of terminal or nonterminal symbols.
 | 
| 28774 | 426  | 
  Then @{text "\<alpha> A\<^sup>(\<^sup>p\<^sup>) \<beta> \<longrightarrow>\<^sub>G \<alpha> \<gamma> \<beta>"} holds if and only if @{text G}
 | 
427  | 
  contains some production @{text "A\<^sup>(\<^sup>q\<^sup>) = \<gamma>"} for @{text "p \<le> q"}.
 | 
|
| 
28769
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
428  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
429  | 
\medskip The following grammar for arithmetic expressions  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
430  | 
demonstrates how binding power and associativity of operators can be  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
431  | 
enforced by priorities.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
432  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
433  | 
  \begin{center}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
434  | 
  \begin{tabular}{rclr}
 | 
| 28774 | 435  | 
  @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim "("} @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim ")"} \\
 | 
| 
28769
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
436  | 
  @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim 0} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
437  | 
  @{text "A\<^sup>(\<^sup>0\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
438  | 
  @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
439  | 
  @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "="} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
440  | 
  \end{tabular}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
441  | 
  \end{center}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
442  | 
  The choice of priorities determines that @{verbatim "-"} binds
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
443  | 
  tighter than @{verbatim "*"}, which binds tighter than @{verbatim
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
444  | 
  "+"}.  Furthermore @{verbatim "+"} associates to the left and
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
445  | 
  @{verbatim "*"} to the right.
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
446  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
447  | 
\medskip For clarity, grammars obey these conventions:  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
448  | 
  \begin{itemize}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
449  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
450  | 
\item All priorities must lie between 0 and 1000.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
451  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
452  | 
\item Priority 0 on the right-hand side and priority 1000 on the  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
453  | 
left-hand side may be omitted.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
454  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
455  | 
  \item The production @{text "A\<^sup>(\<^sup>p\<^sup>) = \<alpha>"} is written as @{text "A = \<alpha>
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
456  | 
(p)"}, i.e.\ the priority of the left-hand side actually appears in  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
457  | 
a column on the far right.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
458  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
459  | 
  \item Alternatives are separated by @{text "|"}.
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
460  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
461  | 
  \item Repetition is indicated by dots @{text "(\<dots>)"} in an informal
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
462  | 
but obvious way.  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
463  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
464  | 
  \end{itemize}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
465  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
466  | 
Using these conventions, the example grammar specification above  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
467  | 
takes the form:  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
468  | 
  \begin{center}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
469  | 
  \begin{tabular}{rclc}
 | 
| 28774 | 470  | 
    @{text A} & @{text "="} & @{verbatim "("} @{text A} @{verbatim ")"} \\
 | 
471  | 
              & @{text "|"} & @{verbatim 0} & \qquad\qquad \\
 | 
|
| 
28769
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
472  | 
              & @{text "|"} & @{text A} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} & @{text "(0)"} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
473  | 
              & @{text "|"} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
474  | 
              & @{text "|"} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
475  | 
  \end{tabular}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
476  | 
  \end{center}
 | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
477  | 
*}  | 
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
478  | 
|
| 
 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 
wenzelm 
parents: 
28767 
diff
changeset
 | 
479  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
480  | 
subsection {* The Pure grammar *}
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
481  | 
|
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
482  | 
text {*
 | 
| 28773 | 483  | 
  The priority grammar of the @{text "Pure"} theory is defined as follows:
 | 
484  | 
||
| 28774 | 485  | 
%FIXME syntax for "index" (?)  | 
486  | 
%FIXME "op" versions of ==> etc. (?)  | 
|
487  | 
||
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
488  | 
  \begin{center}
 | 
| 28773 | 489  | 
  \begin{supertabular}{rclr}
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
490  | 
|
| 28778 | 491  | 
  @{syntax_def (inner) any} & = & @{text "prop  |  logic"} \\\\
 | 
| 28772 | 492  | 
|
| 28778 | 493  | 
  @{syntax_def (inner) prop} & = & @{verbatim "("} @{text prop} @{verbatim ")"} \\
 | 
| 28772 | 494  | 
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\
 | 
495  | 
    & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
 | 
|
| 28773 | 496  | 
    & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{text "\<equiv>"} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
 | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
497  | 
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "&&&"} @{text "prop\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
 | 
| 28772 | 498  | 
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
 | 
| 28773 | 499  | 
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{text "\<Longrightarrow>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
 | 
| 28772 | 500  | 
    & @{text "|"} & @{verbatim "[|"} @{text prop} @{verbatim ";"} @{text "\<dots>"} @{verbatim ";"} @{text prop} @{verbatim "|]"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
 | 
| 28773 | 501  | 
    & @{text "|"} & @{text "\<lbrakk>"} @{text prop} @{verbatim ";"} @{text "\<dots>"} @{verbatim ";"} @{text prop} @{text "\<rbrakk>"} @{text "\<Longrightarrow>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
 | 
| 28772 | 502  | 
    & @{text "|"} & @{verbatim "!!"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\
 | 
| 28773 | 503  | 
    & @{text "|"} & @{text "\<And>"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\
 | 
504  | 
    & @{text "|"} & @{verbatim OFCLASS} @{verbatim "("} @{text type} @{verbatim ","} @{text logic} @{verbatim ")"} \\
 | 
|
505  | 
    & @{text "|"} & @{verbatim SORT_CONSTRAINT} @{verbatim "("} @{text type} @{verbatim ")"} \\
 | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
506  | 
    & @{text "|"} & @{verbatim TERM} @{text logic} \\
 | 
| 28773 | 507  | 
    & @{text "|"} & @{verbatim PROP} @{text aprop} \\\\
 | 
| 28772 | 508  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
509  | 
  @{syntax_def (inner) aprop} & = & @{verbatim "("} @{text aprop} @{verbatim ")"} \\
 | 
| 
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
510  | 
    & @{text "|"} & @{text "id  |  longid  |  var  |  "}@{verbatim "_"}@{text "  |  "}@{verbatim "..."} \\
 | 
| 
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
511  | 
    & @{text "|"} & @{verbatim CONST} @{text "id  |  "}@{verbatim CONST} @{text "longid"} \\
 | 
| 28773 | 512  | 
    & @{text "|"} & @{text "logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)  any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \<dots> any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "(999)"} \\\\
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
513  | 
|
| 28778 | 514  | 
  @{syntax_def (inner) logic} & = & @{verbatim "("} @{text logic} @{verbatim ")"} \\
 | 
| 28772 | 515  | 
    & @{text "|"} & @{text "logic\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\
 | 
| 28773 | 516  | 
    & @{text "|"} & @{text "id  |  longid  |  var  |  "}@{verbatim "_"}@{text "  |  "}@{verbatim "..."} \\
 | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
517  | 
    & @{text "|"} & @{verbatim CONST} @{text "id  |  "}@{verbatim CONST} @{text "longid"} \\
 | 
| 28773 | 518  | 
    & @{text "|"} & @{text "logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)  any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \<dots> any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "(999)"} \\
 | 
| 28772 | 519  | 
    & @{text "|"} & @{verbatim "%"} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\
 | 
| 28773 | 520  | 
    & @{text "|"} & @{text \<lambda>} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\
 | 
| 28772 | 521  | 
    & @{text "|"} & @{verbatim TYPE} @{verbatim "("} @{text type} @{verbatim ")"} \\\\
 | 
522  | 
||
| 28778 | 523  | 
  @{syntax_def (inner) idt} & = & @{verbatim "("} @{text idt} @{verbatim ")"}@{text "  |  id  |  "}@{verbatim "_"} \\
 | 
| 28773 | 524  | 
    & @{text "|"} & @{text id} @{verbatim "::"} @{text type} & @{text "(0)"} \\
 | 
525  | 
    & @{text "|"} & @{verbatim "_"} @{verbatim "::"} @{text type} & @{text "(0)"} \\\\
 | 
|
| 28772 | 526  | 
|
| 28778 | 527  | 
  @{syntax_def (inner) idts} & = & @{text "idt  |  idt\<^sup>(\<^sup>1\<^sup>) idts"} & @{text "(0)"} \\\\
 | 
| 28772 | 528  | 
|
| 28778 | 529  | 
  @{syntax_def (inner) pttrn} & = & @{text idt} \\\\
 | 
| 28772 | 530  | 
|
| 28778 | 531  | 
  @{syntax_def (inner) pttrns} & = & @{text "pttrn  |  pttrn\<^sup>(\<^sup>1\<^sup>) pttrns"} & @{text "(0)"} \\\\
 | 
| 28774 | 532  | 
|
| 28778 | 533  | 
  @{syntax_def (inner) type} & = & @{verbatim "("} @{text type} @{verbatim ")"} \\
 | 
| 28773 | 534  | 
    & @{text "|"} & @{text "tid  |  tvar  |  "}@{verbatim "_"} \\
 | 
535  | 
    & @{text "|"} & @{text "tid"} @{verbatim "::"} @{text "sort  |  tvar  "}@{verbatim "::"} @{text "sort  |  "}@{verbatim "_"} @{verbatim "::"} @{text "sort"} \\
 | 
|
| 28772 | 536  | 
    & @{text "|"} & @{text "id  |  type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) id  |  "}@{verbatim "("} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim ")"} @{text id} \\
 | 
| 29741 | 537  | 
    & @{text "|"} & @{text "longid  |  type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) longid"} \\
 | 
538  | 
    & @{text "|"} & @{verbatim "("} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim ")"} @{text longid} \\
 | 
|
| 28772 | 539  | 
    & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\
 | 
| 28773 | 540  | 
    & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{text "\<Rightarrow>"} @{text type} & @{text "(0)"} \\
 | 
541  | 
    & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim "]"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\
 | 
|
542  | 
    & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim "]"} @{text "\<Rightarrow>"} @{text type} & @{text "(0)"} \\\\
 | 
|
| 28772 | 543  | 
|
| 29741 | 544  | 
  @{syntax_def (inner) sort} & = & @{text "id  |  longid  |  "}@{verbatim "{}"} \\
 | 
545  | 
    & @{text "|"} & @{verbatim "{"} @{text "(id | longid)"} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text "(id | longid)"} @{verbatim "}"} \\
 | 
|
| 28773 | 546  | 
  \end{supertabular}
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
547  | 
  \end{center}
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
548  | 
|
| 28774 | 549  | 
  \medskip Here literal terminals are printed @{verbatim "verbatim"};
 | 
550  | 
  see also \secref{sec:inner-lex} for further token categories of the
 | 
|
551  | 
inner syntax. The meaning of the nonterminals defined by the above  | 
|
552  | 
grammar is as follows:  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
553  | 
|
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
554  | 
  \begin{description}
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
555  | 
|
| 28778 | 556  | 
  \item @{syntax_ref (inner) any} denotes any term.
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
557  | 
|
| 28778 | 558  | 
  \item @{syntax_ref (inner) prop} denotes meta-level propositions,
 | 
559  | 
  which are terms of type @{typ prop}.  The syntax of such formulae of
 | 
|
560  | 
the meta-logic is carefully distinguished from usual conventions for  | 
|
561  | 
  object-logics.  In particular, plain @{text "\<lambda>"}-term notation is
 | 
|
562  | 
  \emph{not} recognized as @{syntax (inner) prop}.
 | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
563  | 
|
| 28778 | 564  | 
  \item @{syntax_ref (inner) aprop} denotes atomic propositions, which
 | 
565  | 
  are embedded into regular @{syntax (inner) prop} by means of an
 | 
|
566  | 
  explicit @{verbatim PROP} token.
 | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
567  | 
|
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
568  | 
  Terms of type @{typ prop} with non-constant head, e.g.\ a plain
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
569  | 
  variable, are printed in this form.  Constants that yield type @{typ
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
570  | 
prop} are expected to provide their own concrete syntax; otherwise  | 
| 28778 | 571  | 
  the printed version will appear like @{syntax (inner) logic} and
 | 
572  | 
  cannot be parsed again as @{syntax (inner) prop}.
 | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
573  | 
|
| 28778 | 574  | 
  \item @{syntax_ref (inner) logic} denotes arbitrary terms of a
 | 
575  | 
  logical type, excluding type @{typ prop}.  This is the main
 | 
|
576  | 
  syntactic category of object-logic entities, covering plain @{text
 | 
|
577  | 
\<lambda>}-term notation (variables, abstraction, application), plus  | 
|
578  | 
anything defined by the user.  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
579  | 
|
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
580  | 
When specifying notation for logical entities, all logical types  | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
581  | 
  (excluding @{typ prop}) are \emph{collapsed} to this single category
 | 
| 28778 | 582  | 
  of @{syntax (inner) logic}.
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
583  | 
|
| 28778 | 584  | 
  \item @{syntax_ref (inner) idt} denotes identifiers, possibly
 | 
585  | 
constrained by types.  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
586  | 
|
| 28778 | 587  | 
  \item @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref
 | 
588  | 
(inner) idt}. This is the most basic category for variables in  | 
|
589  | 
  iterated binders, such as @{text "\<lambda>"} or @{text "\<And>"}.
 | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
590  | 
|
| 28778 | 591  | 
  \item @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns}
 | 
592  | 
denote patterns for abstraction, cases bindings etc. In Pure, these  | 
|
593  | 
  categories start as a merely copy of @{syntax (inner) idt} and
 | 
|
594  | 
  @{syntax (inner) idts}, respectively.  Object-logics may add
 | 
|
595  | 
additional productions for binding forms.  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
596  | 
|
| 28778 | 597  | 
  \item @{syntax_ref (inner) type} denotes types of the meta-logic.
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
598  | 
|
| 28778 | 599  | 
  \item @{syntax_ref (inner) sort} denotes meta-level sorts.
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
600  | 
|
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
601  | 
  \end{description}
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
602  | 
|
| 28774 | 603  | 
Here are some further explanations of certain syntax features.  | 
| 28773 | 604  | 
|
605  | 
  \begin{itemize}
 | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
606  | 
|
| 28778 | 607  | 
  \item In @{syntax (inner) idts}, note that @{text "x :: nat y"} is
 | 
608  | 
  parsed as @{text "x :: (nat y)"}, treating @{text y} like a type
 | 
|
609  | 
  constructor applied to @{text nat}.  To avoid this interpretation,
 | 
|
610  | 
  write @{text "(x :: nat) y"} with explicit parentheses.
 | 
|
| 28773 | 611  | 
|
612  | 
  \item Similarly, @{text "x :: nat y :: nat"} is parsed as @{text "x ::
 | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
613  | 
  (nat y :: nat)"}.  The correct form is @{text "(x :: nat) (y ::
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
614  | 
  nat)"}, or @{text "(x :: nat) y :: nat"} if @{text y} is last in the
 | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
615  | 
sequence of identifiers.  | 
| 28773 | 616  | 
|
617  | 
\item Type constraints for terms bind very weakly. For example,  | 
|
618  | 
  @{text "x < y :: nat"} is normally parsed as @{text "(x < y) ::
 | 
|
619  | 
  nat"}, unless @{text "<"} has a very low priority, in which case the
 | 
|
620  | 
  input is likely to be ambiguous.  The correct form is @{text "x < (y
 | 
|
621  | 
:: nat)"}.  | 
|
622  | 
||
623  | 
\item Constraints may be either written with two literal colons  | 
|
624  | 
  ``@{verbatim "::"}'' or the double-colon symbol @{verbatim "\<Colon>"},
 | 
|
| 28774 | 625  | 
  which actually looks exactly the same in some {\LaTeX} styles.
 | 
| 28773 | 626  | 
|
| 28774 | 627  | 
\item Dummy variables (written as underscore) may occur in different  | 
628  | 
roles.  | 
|
| 28773 | 629  | 
|
630  | 
  \begin{description}
 | 
|
631  | 
||
| 28774 | 632  | 
  \item A type ``@{text "_"}'' or ``@{text "_ :: sort"}'' acts like an
 | 
633  | 
anonymous inference parameter, which is filled-in according to the  | 
|
634  | 
most general type produced by the type-checking phase.  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
635  | 
|
| 28774 | 636  | 
  \item A bound ``@{text "_"}'' refers to a vacuous abstraction, where
 | 
637  | 
the body does not refer to the binding introduced here. As in the  | 
|
638  | 
  term @{term "\<lambda>x _. x"}, which is @{text "\<alpha>"}-equivalent to @{text
 | 
|
639  | 
"\<lambda>x y. x"}.  | 
|
| 28773 | 640  | 
|
| 28774 | 641  | 
  \item A free ``@{text "_"}'' refers to an implicit outer binding.
 | 
642  | 
  Higher definitional packages usually allow forms like @{text "f x _
 | 
|
643  | 
= x"}.  | 
|
| 28773 | 644  | 
|
| 28774 | 645  | 
  \item A schematic ``@{text "_"}'' (within a term pattern, see
 | 
646  | 
  \secref{sec:term-decls}) refers to an anonymous variable that is
 | 
|
647  | 
implicitly abstracted over its context of locally bound variables.  | 
|
648  | 
  For example, this allows pattern matching of @{text "{x. f x = g
 | 
|
649  | 
  x}"} against @{text "{x. _ = _}"}, or even @{text "{_. _ = _}"} by
 | 
|
650  | 
using both bound and schematic dummies.  | 
|
| 28773 | 651  | 
|
652  | 
  \end{description}
 | 
|
653  | 
||
| 28774 | 654  | 
  \item The three literal dots ``@{verbatim "..."}'' may be also
 | 
655  | 
  written as ellipsis symbol @{verbatim "\<dots>"}.  In both cases this
 | 
|
656  | 
refers to a special schematic variable, which is bound in the  | 
|
657  | 
context. This special term abbreviation works nicely with  | 
|
658  | 
  calculational reasoning (\secref{sec:calculation}).
 | 
|
659  | 
||
| 28773 | 660  | 
  \end{itemize}
 | 
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
661  | 
*}  | 
| 
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
662  | 
|
| 28777 | 663  | 
|
| 28774 | 664  | 
section {* Lexical matters \label{sec:inner-lex} *}
 | 
665  | 
||
| 28777 | 666  | 
text {* The inner lexical syntax vaguely resembles the outer one
 | 
667  | 
  (\secref{sec:outer-lex}), but some details are different.  There are
 | 
|
668  | 
two main categories of inner syntax tokens:  | 
|
669  | 
||
670  | 
  \begin{enumerate}
 | 
|
671  | 
||
672  | 
  \item \emph{delimiters} --- the literal tokens occurring in
 | 
|
673  | 
productions of the given priority grammar (cf.\  | 
|
674  | 
  \secref{sec:priority-grammar});
 | 
|
675  | 
||
676  | 
  \item \emph{named tokens} --- various categories of identifiers etc.
 | 
|
677  | 
||
678  | 
  \end{enumerate}
 | 
|
679  | 
||
680  | 
Delimiters override named tokens and may thus render certain  | 
|
681  | 
identifiers inaccessible. Sometimes the logical context admits  | 
|
682  | 
alternative ways to refer to the same entity, potentially via  | 
|
683  | 
qualified names.  | 
|
684  | 
||
685  | 
\medskip The categories for named tokens are defined once and for  | 
|
686  | 
all as follows, reusing some categories of the outer token syntax  | 
|
687  | 
  (\secref{sec:outer-lex}).
 | 
|
688  | 
||
689  | 
  \begin{center}
 | 
|
690  | 
  \begin{supertabular}{rcl}
 | 
|
691  | 
    @{syntax_def (inner) id} & = & @{syntax_ref ident} \\
 | 
|
692  | 
    @{syntax_def (inner) longid} & = & @{syntax_ref longident} \\
 | 
|
693  | 
    @{syntax_def (inner) var} & = & @{syntax_ref var} \\
 | 
|
694  | 
    @{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\
 | 
|
695  | 
    @{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\
 | 
|
696  | 
    @{syntax_def (inner) num} & = & @{syntax_ref nat}@{text "  |  "}@{verbatim "-"}@{syntax_ref nat} \\
 | 
|
| 29157 | 697  | 
    @{syntax_def (inner) float_token} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text "  |  "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\
 | 
| 28777 | 698  | 
    @{syntax_def (inner) xnum} & = & @{verbatim "#"}@{syntax_ref nat}@{text "  |  "}@{verbatim "#-"}@{syntax_ref nat} \\
 | 
699  | 
||
700  | 
    @{syntax_def (inner) xstr} & = & @{verbatim "''"} @{text "\<dots>"} @{verbatim "''"} \\
 | 
|
701  | 
  \end{supertabular}
 | 
|
702  | 
  \end{center}
 | 
|
703  | 
||
| 29157 | 704  | 
  The token categories @{syntax (inner) num}, @{syntax (inner)
 | 
705  | 
  float_token}, @{syntax (inner) xnum}, and @{syntax (inner) xstr} are
 | 
|
706  | 
not used in Pure. Object-logics may implement numerals and string  | 
|
707  | 
constants by adding appropriate syntax declarations, together with  | 
|
708  | 
some translation functions (e.g.\ see Isabelle/HOL).  | 
|
709  | 
||
710  | 
  The derived categories @{syntax_def (inner) num_const} and
 | 
|
711  | 
  @{syntax_def (inner) float_const} provide robust access to @{syntax
 | 
|
712  | 
  (inner) num}, and @{syntax (inner) float_token}, respectively: the
 | 
|
713  | 
syntax tree holds a syntactic constant instead of a free variable.  | 
|
| 28777 | 714  | 
*}  | 
| 28774 | 715  | 
|
| 
28770
 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 
wenzelm 
parents: 
28769 
diff
changeset
 | 
716  | 
|
| 28762 | 717  | 
section {* Syntax and translations \label{sec:syn-trans} *}
 | 
718  | 
||
719  | 
text {*
 | 
|
720  | 
  \begin{matharray}{rcl}
 | 
|
| 
41229
 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
721  | 
    @{command_def "nonterminal"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 28762 | 722  | 
    @{command_def "syntax"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
723  | 
    @{command_def "no_syntax"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
724  | 
    @{command_def "translations"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
725  | 
    @{command_def "no_translations"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
726  | 
  \end{matharray}
 | 
|
727  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
728  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
729  | 
    @@{command nonterminal} (@{syntax name} + @'and')
 | 
| 28762 | 730  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
731  | 
    (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (@{syntax constdecl} +)
 | 
| 28762 | 732  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
733  | 
    (@@{command translations} | @@{command no_translations})
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
734  | 
      (transpat ('==' | '=>' | '<=' | '\<rightleftharpoons>' | '\<rightharpoonup>' | '\<leftharpoondown>') transpat +)
 | 
| 28762 | 735  | 
;  | 
736  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
737  | 
    mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')')
 | 
| 28762 | 738  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
739  | 
    transpat: ('(' @{syntax nameref} ')')? @{syntax string}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
740  | 
"}  | 
| 28762 | 741  | 
|
742  | 
  \begin{description}
 | 
|
743  | 
||
| 
41229
 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
744  | 
  \item @{command "nonterminal"}~@{text c} declares a type
 | 
| 28762 | 745  | 
  constructor @{text c} (without arguments) to act as purely syntactic
 | 
746  | 
type: a nonterminal symbol of the inner syntax.  | 
|
747  | 
||
748  | 
  \item @{command "syntax"}~@{text "(mode) decls"} is similar to
 | 
|
749  | 
  @{command "consts"}~@{text decls}, except that the actual logical
 | 
|
750  | 
signature extension is omitted. Thus the context free grammar of  | 
|
751  | 
Isabelle's inner syntax may be augmented in arbitrary ways,  | 
|
752  | 
  independently of the logic.  The @{text mode} argument refers to the
 | 
|
753  | 
  print mode that the grammar rules belong; unless the @{keyword_ref
 | 
|
754  | 
"output"} indicator is given, all productions are added both to the  | 
|
755  | 
input and output grammar.  | 
|
756  | 
||
757  | 
  \item @{command "no_syntax"}~@{text "(mode) decls"} removes grammar
 | 
|
758  | 
  declarations (and translations) resulting from @{text decls}, which
 | 
|
759  | 
  are interpreted in the same manner as for @{command "syntax"} above.
 | 
|
760  | 
||
761  | 
  \item @{command "translations"}~@{text rules} specifies syntactic
 | 
|
762  | 
  translation rules (i.e.\ macros): parse~/ print rules (@{text "\<rightleftharpoons>"}),
 | 
|
763  | 
  parse rules (@{text "\<rightharpoonup>"}), or print rules (@{text "\<leftharpoondown>"}).
 | 
|
764  | 
Translation patterns may be prefixed by the syntactic category to be  | 
|
765  | 
  used for parsing; the default is @{text logic}.
 | 
|
766  | 
||
767  | 
  \item @{command "no_translations"}~@{text rules} removes syntactic
 | 
|
768  | 
translation rules, which are interpreted in the same manner as for  | 
|
769  | 
  @{command "translations"} above.
 | 
|
770  | 
||
771  | 
  \end{description}
 | 
|
772  | 
*}  | 
|
773  | 
||
774  | 
||
| 
28779
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
775  | 
section {* Syntax translation functions \label{sec:tr-funs} *}
 | 
| 28762 | 776  | 
|
777  | 
text {*
 | 
|
778  | 
  \begin{matharray}{rcl}
 | 
|
779  | 
    @{command_def "parse_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
780  | 
    @{command_def "parse_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
781  | 
    @{command_def "print_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
782  | 
    @{command_def "typed_print_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
783  | 
    @{command_def "print_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
|
784  | 
  \end{matharray}
 | 
|
785  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
786  | 
  @{rail "
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
787  | 
  ( @@{command parse_ast_translation} | @@{command parse_translation} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
788  | 
    @@{command print_translation} | @@{command typed_print_translation} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
789  | 
    @@{command print_ast_translation}) ('(' @'advanced' ')')? @{syntax text}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
790  | 
"}  | 
| 28762 | 791  | 
|
792  | 
Syntax translation functions written in ML admit almost arbitrary  | 
|
793  | 
manipulations of Isabelle's inner syntax. Any of the above commands  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
42358 
diff
changeset
 | 
794  | 
  have a single @{syntax text} argument that refers to an ML
 | 
| 28762 | 795  | 
expression of appropriate type, which are as follows by default:  | 
796  | 
||
797  | 
%FIXME proper antiquotations  | 
|
798  | 
\begin{ttbox}
 | 
|
799  | 
val parse_ast_translation : (string * (ast list -> ast)) list  | 
|
800  | 
val parse_translation : (string * (term list -> term)) list  | 
|
801  | 
val print_translation : (string * (term list -> term)) list  | 
|
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
802  | 
val typed_print_translation : (string * (typ -> term list -> term)) list  | 
| 28762 | 803  | 
val print_ast_translation : (string * (ast list -> ast)) list  | 
804  | 
\end{ttbox}
 | 
|
805  | 
||
806  | 
  If the @{text "(advanced)"} option is given, the corresponding
 | 
|
807  | 
translation functions may depend on the current theory or proof  | 
|
808  | 
context. This allows to implement advanced syntax mechanisms, as  | 
|
809  | 
translations functions may refer to specific theory declarations or  | 
|
810  | 
auxiliary proof data.  | 
|
811  | 
||
| 30397 | 812  | 
  See also \cite{isabelle-ref} for more information on the general
 | 
813  | 
concept of syntax transformations in Isabelle.  | 
|
| 28762 | 814  | 
|
815  | 
%FIXME proper antiquotations  | 
|
816  | 
\begin{ttbox}
 | 
|
817  | 
val parse_ast_translation:  | 
|
818  | 
(string * (Proof.context -> ast list -> ast)) list  | 
|
819  | 
val parse_translation:  | 
|
820  | 
(string * (Proof.context -> term list -> term)) list  | 
|
821  | 
val print_translation:  | 
|
822  | 
(string * (Proof.context -> term list -> term)) list  | 
|
823  | 
val typed_print_translation:  | 
|
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
824  | 
(string * (Proof.context -> typ -> term list -> term)) list  | 
| 28762 | 825  | 
val print_ast_translation:  | 
826  | 
(string * (Proof.context -> ast list -> ast)) list  | 
|
827  | 
\end{ttbox}
 | 
|
828  | 
*}  | 
|
829  | 
||
| 
28779
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
830  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
831  | 
section {* Inspecting the syntax *}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
832  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
833  | 
text {*
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
834  | 
  \begin{matharray}{rcl}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
835  | 
    @{command_def "print_syntax"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
836  | 
  \end{matharray}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
837  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
838  | 
  \begin{description}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
839  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
840  | 
  \item @{command "print_syntax"} prints the inner syntax of the
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
841  | 
current context. The output can be quite large; the most important  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
842  | 
sections are explained below.  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
843  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
844  | 
  \begin{description}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
845  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
846  | 
  \item @{text "lexicon"} lists the delimiters of the inner token
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
847  | 
  language; see \secref{sec:inner-lex}.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
848  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
849  | 
  \item @{text "prods"} lists the productions of the underlying
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
850  | 
  priority grammar; see \secref{sec:priority-grammar}.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
851  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
852  | 
  The nonterminal @{text "A\<^sup>(\<^sup>p\<^sup>)"} is rendered in plain text as @{text
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
853  | 
"A[p]"}; delimiters are quoted. Many productions have an extra  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
854  | 
  @{text "\<dots> => name"}.  These names later become the heads of parse
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
855  | 
trees; they also guide the pretty printer.  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
856  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
857  | 
  Productions without such parse tree names are called \emph{copy
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
858  | 
productions}. Their right-hand side must have exactly one  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
859  | 
nonterminal symbol (or named token). The parser does not create a  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
860  | 
new parse tree node for copy productions, but simply returns the  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
861  | 
parse tree of the right-hand symbol.  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
862  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
863  | 
If the right-hand side of a copy production consists of a single  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
864  | 
  nonterminal without any delimiters, then it is called a \emph{chain
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
865  | 
production}. Chain productions act as abbreviations: conceptually,  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
866  | 
they are removed from the grammar by adding new productions.  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
867  | 
Priority information attached to chain productions is ignored; only  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
868  | 
  the dummy value @{text "-1"} is displayed.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
869  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28779 
diff
changeset
 | 
870  | 
  \item @{text "print modes"} lists the alternative print modes
 | 
| 
28779
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
871  | 
  provided by this grammar; see \secref{sec:print-modes}.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
872  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
873  | 
  \item @{text "parse_rules"} and @{text "print_rules"} relate to
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
874  | 
  syntax translations (macros); see \secref{sec:syn-trans}.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
875  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
876  | 
  \item @{text "parse_ast_translation"} and @{text
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
877  | 
"print_ast_translation"} list sets of constants that invoke  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
878  | 
translation functions for abstract syntax trees, which are only  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
879  | 
  required in very special situations; see \secref{sec:tr-funs}.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
880  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
881  | 
  \item @{text "parse_translation"} and @{text "print_translation"}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
882  | 
list the sets of constants that invoke regular translation  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
883  | 
  functions; see \secref{sec:tr-funs}.
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
884  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
885  | 
  \end{description}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
886  | 
|
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
887  | 
  \end{description}
 | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
888  | 
*}  | 
| 
 
698960f08652
separate section "Inspecting the syntax" for print_syntax;
 
wenzelm 
parents: 
28778 
diff
changeset
 | 
889  | 
|
| 28762 | 890  | 
end  |