11647
|
1 |
(*<*)
|
|
2 |
theory Documents = Main:
|
|
3 |
(*>*)
|
|
4 |
|
12629
|
5 |
section {* Concrete syntax \label{sec:concrete-syntax} *}
|
|
6 |
|
|
7 |
text {*
|
|
8 |
Concerning Isabelle's ``inner'' language of simply-typed @{text
|
|
9 |
\<lambda>}-calculus, the core concept of Isabelle's elaborate infrastructure
|
|
10 |
for concrete syntax is that of general \emph{mixfix
|
|
11 |
annotations}\index{mixfix annotations|bold}. Associated with any
|
|
12 |
kind of name and type declaration, mixfixes give rise both to
|
|
13 |
grammar productions for the parser and output templates for the
|
|
14 |
pretty printer.
|
|
15 |
|
|
16 |
In full generality, the whole affair of parser and pretty printer
|
|
17 |
configuration is rather subtle. Any syntax specifications given by
|
|
18 |
end-users need to interact properly with the existing setup of
|
|
19 |
Isabelle/Pure and Isabelle/HOL; see \cite{isabelle-ref} for further
|
|
20 |
details. It is particularly important to get the precedence of new
|
|
21 |
syntactic constructs right, avoiding ambiguities with existing
|
|
22 |
elements.
|
|
23 |
|
|
24 |
\medskip Subsequently we introduce a few simple declaration forms
|
|
25 |
that already cover the most common situations fairly well.
|
|
26 |
*}
|
|
27 |
|
|
28 |
|
|
29 |
subsection {* Infixes *}
|
|
30 |
|
|
31 |
text {*
|
|
32 |
Syntax annotations may be included wherever constants are declared
|
|
33 |
directly or indirectly, including \isacommand{consts},
|
|
34 |
\isacommand{constdefs}, or \isacommand{datatype} (for the
|
|
35 |
constructor operations). Type-constructors may be annotated as
|
|
36 |
well, although this is less frequently encountered in practice
|
|
37 |
(@{text "*"} and @{text "+"} types may come to mind).
|
|
38 |
|
|
39 |
Infix declarations\index{infix annotations|bold} provide a useful
|
|
40 |
special case of mixfixes, where users need not care about the full
|
|
41 |
details of priorities, nesting, spacing, etc. The subsequent
|
|
42 |
example of the exclusive-or operation on boolean values illustrates
|
|
43 |
typical infix declarations.
|
|
44 |
*}
|
|
45 |
|
|
46 |
constdefs
|
|
47 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]" 60)
|
|
48 |
"A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
|
|
49 |
|
|
50 |
text {*
|
|
51 |
Any curried function with at least two arguments may be associated
|
|
52 |
with infix syntax: @{text "xor A B"} and @{text "A [+] B"} refer to
|
|
53 |
the same expression internally. In partial applications with less
|
|
54 |
than two operands there is a special notation with \isa{op} prefix:
|
|
55 |
@{text xor} without arguments is represented as @{text "op [+]"};
|
|
56 |
combined with plain prefix application this turns @{text "xor A"}
|
|
57 |
into @{text "op [+] A"}.
|
|
58 |
|
|
59 |
\medskip The string @{text [source] "[+]"} in the above declaration
|
|
60 |
refers to the bit of concrete syntax to represent the operator,
|
|
61 |
while the number @{text 60} determines the precedence of the whole
|
|
62 |
construct.
|
|
63 |
|
|
64 |
As it happens, Isabelle/HOL already spends many popular combinations
|
|
65 |
of ASCII symbols for its own use, including both @{text "+"} and
|
|
66 |
@{text "++"}. Slightly more awkward combinations like the present
|
|
67 |
@{text "[+]"} tend to be available for user extensions. The current
|
|
68 |
arrangement of inner syntax may be inspected via
|
|
69 |
\commdx{print\protect\_syntax}, albeit its output is enormous.
|
|
70 |
|
|
71 |
Operator precedence also needs some special considerations. The
|
|
72 |
admissible range is 0--1000. Very low or high priorities are
|
|
73 |
basically reserved for the meta-logic. Syntax of Isabelle/HOL
|
|
74 |
mainly uses the range of 10--100: the equality infix @{text "="} is
|
|
75 |
centered at 50, logical connectives (like @{text "\<or>"} and @{text
|
|
76 |
"\<and>"}) are below 50, and algebraic ones (like @{text "+"} and @{text
|
|
77 |
"*"}) above 50. User syntax should strive to coexist with common
|
|
78 |
HOL forms, or use the mostly unused range 100--900.
|
|
79 |
|
|
80 |
\medskip The keyword \isakeyword{infixl} specifies an operator that
|
|
81 |
is nested to the \emph{left}: in iterated applications the more
|
|
82 |
complex expression appears on the left-hand side: @{term "A [+] B
|
|
83 |
[+] C"} stands for @{text "(A [+] B) [+] C"}. Similarly,
|
|
84 |
\isakeyword{infixr} refers to nesting to the \emph{right}, which
|
|
85 |
would turn @{term "A [+] B [+] C"} into @{text "A [+] (B [+] C)"}.
|
|
86 |
In contrast, a \emph{non-oriented} declaration via
|
|
87 |
\isakeyword{infix} would always demand explicit parentheses.
|
|
88 |
|
|
89 |
Many binary operations observe the associative law, so the exact
|
|
90 |
grouping does not matter. Nevertheless, formal statements need be
|
|
91 |
given in a particular format, associativity needs to be treated
|
|
92 |
explicitly within the logic. Exclusive-or is happens to be
|
|
93 |
associative, as shown below.
|
|
94 |
*}
|
|
95 |
|
|
96 |
lemma xor_assoc: "(A [+] B) [+] C = A [+] (B [+] C)"
|
|
97 |
by (auto simp add: xor_def)
|
|
98 |
|
|
99 |
text {*
|
|
100 |
Such rules may be used in simplification to regroup nested
|
|
101 |
expressions as required. Note that the system would actually print
|
|
102 |
the above statement as @{term "A [+] B [+] C = A [+] (B [+] C)"}
|
|
103 |
(due to nesting to the left). We have preferred to give the fully
|
|
104 |
parenthesized form in the text for clarity.
|
|
105 |
*}
|
|
106 |
|
|
107 |
(*<*)(*FIXME*)
|
|
108 |
subsection {* Mathematical symbols \label{sec:thy-present-symbols} *}
|
|
109 |
|
|
110 |
text {*
|
|
111 |
The limitations of the ASCII character set pose a serious
|
|
112 |
limitations for representing mathematical notation. Even worse
|
|
113 |
many handsome combinations have already been used up by HOL
|
|
114 |
itself. Luckily, Isabelle supports infinitely many \emph{named
|
|
115 |
symbols}. FIXME
|
|
116 |
|
|
117 |
*}
|
|
118 |
|
|
119 |
(*<*)
|
|
120 |
hide const xor
|
|
121 |
ML_setup {* Context.>> (Theory.add_path "1") *}
|
|
122 |
(*>*)
|
|
123 |
constdefs
|
|
124 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>" 60)
|
|
125 |
"A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
|
|
126 |
|
|
127 |
(*<*)
|
|
128 |
local
|
|
129 |
(*>*)
|
|
130 |
|
|
131 |
|
|
132 |
subsection {* Prefixes *}
|
|
133 |
|
|
134 |
text {*
|
|
135 |
Prefix declarations are an even more degenerate form of mixfix
|
|
136 |
annotation, which allow a arbitrary symbolic token to be used for FIXME
|
|
137 |
*}
|
|
138 |
|
|
139 |
datatype currency =
|
|
140 |
Euro nat ("\<euro>")
|
|
141 |
| Pounds nat ("\<pounds>")
|
|
142 |
| Yen nat ("\<yen>")
|
|
143 |
| Dollar nat ("$")
|
|
144 |
|
|
145 |
text {*
|
|
146 |
FIXME The predefined Isabelle symbols used above are \verb,\<euro>,,
|
|
147 |
\verb,\<pounds>,, \verb,\<yen>,, and \verb,\<dollar>,.
|
|
148 |
|
|
149 |
The above syntax annotation makes \verb,\<euro>, stand for the
|
|
150 |
datatype constructor constant @{text Euro}, which happens to be a
|
|
151 |
function @{typ "nat \<Rightarrow> currency"}. Using plain application syntax
|
|
152 |
we may write @{text "Euro 10"} as @{term "\<euro> 10"}. So we already
|
|
153 |
achieve a decent syntactic representation without having to consider
|
|
154 |
arguments and precedences of general mixfix annotations -- prefixes
|
|
155 |
are already sufficient.
|
|
156 |
*}
|
|
157 |
|
|
158 |
|
|
159 |
subsection {* Syntax translations \label{sec:def-translations} *}
|
|
160 |
|
|
161 |
text{*
|
|
162 |
FIXME
|
|
163 |
|
|
164 |
\index{syntax translations|(}%
|
|
165 |
\index{translations@\isacommand {translations} (command)|(}
|
|
166 |
Isabelle offers an additional definitional facility,
|
|
167 |
\textbf{syntax translations}.
|
|
168 |
They resemble macros: upon parsing, the defined concept is immediately
|
|
169 |
replaced by its definition. This effect is reversed upon printing. For example,
|
|
170 |
the symbol @{text"\<noteq>"} is defined via a syntax translation:
|
|
171 |
*}
|
|
172 |
|
|
173 |
translations "x \<noteq> y" \<rightleftharpoons> "\<not>(x = y)"
|
|
174 |
|
|
175 |
text{*\index{$IsaEqTrans@\isasymrightleftharpoons}
|
|
176 |
\noindent
|
|
177 |
Internally, @{text"\<noteq>"} never appears.
|
|
178 |
|
|
179 |
In addition to @{text"\<rightleftharpoons>"} there are
|
|
180 |
@{text"\<rightharpoonup>"}\index{$IsaEqTrans1@\isasymrightharpoonup}
|
|
181 |
and @{text"\<leftharpoondown>"}\index{$IsaEqTrans2@\isasymleftharpoondown}
|
|
182 |
for uni-directional translations, which only affect
|
|
183 |
parsing or printing. This tutorial will not cover the details of
|
|
184 |
translations. We have mentioned the concept merely because it
|
|
185 |
crops up occasionally; a number of HOL's built-in constructs are defined
|
|
186 |
via translations. Translations are preferable to definitions when the new
|
|
187 |
concept is a trivial variation on an existing one. For example, we
|
|
188 |
don't need to derive new theorems about @{text"\<noteq>"}, since existing theorems
|
|
189 |
about @{text"="} still apply.%
|
|
190 |
\index{syntax translations|)}%
|
|
191 |
\index{translations@\isacommand {translations} (command)|)}
|
|
192 |
*}
|
|
193 |
|
|
194 |
|
|
195 |
section {* Document preparation *}
|
|
196 |
|
|
197 |
subsection {* Batch-mode sessions *}
|
|
198 |
|
|
199 |
subsection {* {\LaTeX} macros *}
|
|
200 |
|
|
201 |
subsubsection {* Structure markup *}
|
|
202 |
|
|
203 |
subsubsection {* Symbols and characters *}
|
|
204 |
|
|
205 |
text {*
|
|
206 |
FIXME
|
|
207 |
|
|
208 |
|
|
209 |
*}
|
|
210 |
|
11647
|
211 |
(*<*)
|
|
212 |
end
|
|
213 |
(*>*)
|
12629
|
214 |
(*>*)
|