| author | blanchet | 
| Thu, 28 Jul 2011 11:43:45 +0200 | |
| changeset 43996 | 4d1270ddf042 | 
| parent 42936 | 48a0a9db3453 | 
| child 45600 | 1bbbac9a0cb0 | 
| permissions | -rw-r--r-- | 
| 26869 | 1  | 
%  | 
2  | 
\begin{isabellebody}%
 | 
|
3  | 
\def\isabellecontext{Spec}%
 | 
|
4  | 
%  | 
|
5  | 
\isadelimtheory  | 
|
6  | 
%  | 
|
7  | 
\endisadelimtheory  | 
|
8  | 
%  | 
|
9  | 
\isatagtheory  | 
|
10  | 
\isacommand{theory}\isamarkupfalse%
 | 
|
11  | 
\ Spec\isanewline  | 
|
| 42651 | 12  | 
\isakeyword{imports}\ Base\ Main\isanewline
 | 
| 26869 | 13  | 
\isakeyword{begin}%
 | 
14  | 
\endisatagtheory  | 
|
15  | 
{\isafoldtheory}%
 | 
|
16  | 
%  | 
|
17  | 
\isadelimtheory  | 
|
18  | 
%  | 
|
19  | 
\endisadelimtheory  | 
|
20  | 
%  | 
|
| 42936 | 21  | 
\isamarkupchapter{Specifications%
 | 
| 26869 | 22  | 
}  | 
23  | 
\isamarkuptrue%  | 
|
24  | 
%  | 
|
| 29746 | 25  | 
\begin{isamarkuptext}%
 | 
26  | 
The Isabelle/Isar theory format integrates specifications and  | 
|
27  | 
proofs, supporting interactive development with unlimited undo  | 
|
28  | 
operation. There is an integrated document preparation system (see  | 
|
29  | 
  \chref{ch:document-prep}), for typesetting formal developments
 | 
|
30  | 
together with informal text. The resulting hyper-linked PDF  | 
|
31  | 
documents can be used both for WWW presentation and printed copies.  | 
|
32  | 
||
33  | 
  The Isar proof language (see \chref{ch:proofs}) is embedded into the
 | 
|
34  | 
theory language as a proper sub-language. Proof mode is entered by  | 
|
35  | 
  stating some \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}} or \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}} at the theory
 | 
|
36  | 
  level, and left again with the final conclusion (e.g.\ via \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}).  Some theory specification mechanisms also require a proof,
 | 
|
37  | 
  such as \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} in HOL, which demands non-emptiness of
 | 
|
38  | 
the representing sets.%  | 
|
39  | 
\end{isamarkuptext}%
 | 
|
40  | 
\isamarkuptrue%  | 
|
41  | 
%  | 
|
| 26870 | 42  | 
\isamarkupsection{Defining theories \label{sec:begin-thy}%
 | 
43  | 
}  | 
|
44  | 
\isamarkuptrue%  | 
|
45  | 
%  | 
|
46  | 
\begin{isamarkuptext}%
 | 
|
47  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 48  | 
    \indexdef{}{command}{theory}\hypertarget{command.theory}{\hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}toplevel\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
49  | 
    \indexdef{global}{command}{end}\hypertarget{command.global.end}{\hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ toplevel{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 26870 | 50  | 
  \end{matharray}
 | 
51  | 
||
| 28788 | 52  | 
Isabelle/Isar theories are defined via theory files, which may  | 
53  | 
contain both specifications and proofs; occasionally definitional  | 
|
54  | 
mechanisms also require some explicit proof. The theory body may be  | 
|
55  | 
  sub-structured by means of \emph{local theory targets}, such as
 | 
|
56  | 
  \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} and \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}.
 | 
|
| 26870 | 57  | 
|
| 28788 | 58  | 
  The first proper command of a theory is \hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}, which
 | 
59  | 
indicates imports of previous theories and optional dependencies on  | 
|
60  | 
other source files (usually in ML). Just preceding the initial  | 
|
61  | 
  \hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}} command there may be an optional \hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}} declaration, which is only relevant to document
 | 
|
62  | 
preparation: see also the other section markup commands in  | 
|
63  | 
  \secref{sec:markup}.
 | 
|
64  | 
||
65  | 
  A theory is concluded by a final \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}} command,
 | 
|
66  | 
one that does not belong to a local theory target. No further  | 
|
67  | 
  commands may follow such a global \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}},
 | 
|
68  | 
although some user-interfaces might pretend that trailing input is  | 
|
69  | 
admissible.  | 
|
| 26870 | 70  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
71  | 
  \begin{railoutput}
 | 
| 42704 | 72  | 
\rail@begin{4}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
73  | 
\rail@term{\hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
74  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 42704 | 75  | 
\rail@cr{2}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
76  | 
\rail@term{\isa{\isakeyword{imports}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
77  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
78  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 42704 | 79  | 
\rail@nextplus{3}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
80  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
81  | 
\rail@bar  | 
| 42704 | 82  | 
\rail@nextbar{3}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
83  | 
\rail@nont{\isa{uses}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
84  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
85  | 
\rail@term{\isa{\isakeyword{begin}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
86  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
87  | 
\rail@begin{3}{\isa{uses}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
88  | 
\rail@term{\isa{\isakeyword{uses}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
89  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
90  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
91  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
92  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
93  | 
\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
94  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
95  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
96  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
97  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
98  | 
\end{railoutput}
 | 
| 26870 | 99  | 
|
100  | 
||
| 28788 | 101  | 
  \begin{description}
 | 
| 26870 | 102  | 
|
| 40406 | 103  | 
  \item \hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C494D504F5254533E}{\isasymIMPORTS}}\ B\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ B\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 28788 | 104  | 
  starts a new theory \isa{A} based on the merge of existing
 | 
| 40406 | 105  | 
  theories \isa{{\isaliteral{22}{\isachardoublequote}}B\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ B\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}.
 | 
| 26870 | 106  | 
|
| 28788 | 107  | 
Due to the possibility to import more than one ancestor, the  | 
108  | 
resulting theory structure of an Isabelle session forms a directed  | 
|
109  | 
acyclic graph (DAG). Isabelle's theory loader ensures that the  | 
|
110  | 
sources contributing to the development graph are always up-to-date:  | 
|
111  | 
changed files are automatically reloaded whenever a theory header  | 
|
112  | 
specification is processed.  | 
|
| 26870 | 113  | 
|
| 26902 | 114  | 
  The optional \indexdef{}{keyword}{uses}\hypertarget{keyword.uses}{\hyperlink{keyword.uses}{\mbox{\isa{\isakeyword{uses}}}}} specification declares additional
 | 
| 26870 | 115  | 
dependencies on extra files (usually ML sources). Files will be  | 
| 28788 | 116  | 
loaded immediately (as ML), unless the name is parenthesized. The  | 
117  | 
latter case records a dependency that needs to be resolved later in  | 
|
118  | 
  the text, usually via explicit \indexref{}{command}{use}\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}} for ML files;
 | 
|
119  | 
other file formats require specific load commands defined by the  | 
|
120  | 
corresponding tools or packages.  | 
|
| 26870 | 121  | 
|
| 28788 | 122  | 
  \item \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}} concludes the current theory
 | 
123  | 
definition. Note that local theory targets involve a local  | 
|
124  | 
  \hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}}, which is clear from the nesting.
 | 
|
| 27042 | 125  | 
|
| 28788 | 126  | 
  \end{description}%
 | 
| 27042 | 127  | 
\end{isamarkuptext}%
 | 
128  | 
\isamarkuptrue%  | 
|
129  | 
%  | 
|
130  | 
\isamarkupsection{Local theory targets \label{sec:target}%
 | 
|
131  | 
}  | 
|
132  | 
\isamarkuptrue%  | 
|
133  | 
%  | 
|
134  | 
\begin{isamarkuptext}%
 | 
|
135  | 
A local theory target is a context managed separately within the  | 
|
136  | 
enclosing theory. Contexts may introduce parameters (fixed  | 
|
137  | 
variables) and assumptions (hypotheses). Definitions and theorems  | 
|
138  | 
depending on the context may be added incrementally later on. Named  | 
|
139  | 
  contexts refer to locales (cf.\ \secref{sec:locale}) or type classes
 | 
|
| 40406 | 140  | 
  (cf.\ \secref{sec:class}); the name ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' signifies the
 | 
| 27042 | 141  | 
global theory context.  | 
142  | 
||
143  | 
  \begin{matharray}{rcll}
 | 
|
| 40406 | 144  | 
    \indexdef{}{command}{context}\hypertarget{command.context}{\hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
145  | 
    \indexdef{local}{command}{end}\hypertarget{command.local.end}{\hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27042 | 146  | 
  \end{matharray}
 | 
147  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
148  | 
  \begin{railoutput}
 | 
| 42662 | 149  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
150  | 
\rail@term{\hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
151  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
152  | 
\rail@term{\isa{\isakeyword{begin}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
153  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
154  | 
\rail@begin{1}{\indexdef{}{syntax}{target}\hypertarget{syntax.target}{\hyperlink{syntax.target}{\mbox{\isa{target}}}}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
155  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
156  | 
\rail@term{\isa{\isakeyword{in}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
157  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
158  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
159  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
160  | 
\end{railoutput}
 | 
| 27042 | 161  | 
|
162  | 
||
| 28788 | 163  | 
  \begin{description}
 | 
| 27042 | 164  | 
|
| 40406 | 165  | 
  \item \hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} recommences an
 | 
| 27042 | 166  | 
  existing locale or class context \isa{c}.  Note that locale and
 | 
| 27052 | 167  | 
  class definitions allow to include the \hyperlink{keyword.begin}{\mbox{\isa{\isakeyword{begin}}}} keyword as
 | 
168  | 
well, in order to continue the local theory immediately after the  | 
|
169  | 
initial specification.  | 
|
| 27042 | 170  | 
|
| 28788 | 171  | 
  \item \hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}} concludes the current local theory
 | 
| 27042 | 172  | 
and continues the enclosing global theory. Note that a global  | 
173  | 
  \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}} has a different meaning: it concludes the
 | 
|
174  | 
  theory itself (\secref{sec:begin-thy}).
 | 
|
175  | 
||
| 40406 | 176  | 
  \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}}\indexdef{}{keyword}{in}\hypertarget{keyword.in}{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} given after any
 | 
| 29746 | 177  | 
local theory command specifies an immediate target, e.g.\  | 
| 40406 | 178  | 
  ``\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C494E3E}{\isasymIN}}\ c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' or ``\hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C494E3E}{\isasymIN}}\ c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}''.  This works both in a local or
 | 
| 27042 | 179  | 
global theory context; the current target context will be suspended  | 
| 40406 | 180  | 
  for this command only.  Note that ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C494E3E}{\isasymIN}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' will
 | 
| 27042 | 181  | 
always produce a global result independently of the current target  | 
182  | 
context.  | 
|
183  | 
||
| 28788 | 184  | 
  \end{description}
 | 
| 27042 | 185  | 
|
186  | 
The exact meaning of results produced within a local theory context  | 
|
187  | 
depends on the underlying target infrastructure (locale, type class  | 
|
188  | 
etc.). The general idea is as follows, considering a context named  | 
|
| 40406 | 189  | 
  \isa{c} with parameter \isa{x} and assumption \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
| 27042 | 190  | 
|
191  | 
Definitions are exported by introducing a global version with  | 
|
192  | 
additional arguments; a syntactic abbreviation links the long form  | 
|
193  | 
with the abstract version of the target context. For example,  | 
|
| 40406 | 194  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} becomes \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{2E}{\isachardot}}a\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} at the theory
 | 
195  | 
  level (for arbitrary \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{22}{\isachardoublequote}}}), together with a local
 | 
|
196  | 
  abbreviation \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ c{\isaliteral{2E}{\isachardot}}a\ x{\isaliteral{22}{\isachardoublequote}}} in the target context (for the
 | 
|
| 27042 | 197  | 
  fixed parameter \isa{x}).
 | 
198  | 
||
199  | 
Theorems are exported by discharging the assumptions and  | 
|
| 40406 | 200  | 
  generalizing the parameters of the context.  For example, \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} becomes \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{2E}{\isachardot}}a{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}, again for arbitrary
 | 
201  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{22}{\isachardoublequote}}}.%
 | 
|
| 27042 | 202  | 
\end{isamarkuptext}%
 | 
203  | 
\isamarkuptrue%  | 
|
204  | 
%  | 
|
205  | 
\isamarkupsection{Basic specification elements%
 | 
|
206  | 
}  | 
|
207  | 
\isamarkuptrue%  | 
|
208  | 
%  | 
|
209  | 
\begin{isamarkuptext}%
 | 
|
210  | 
\begin{matharray}{rcll}
 | 
|
| 40406 | 211  | 
    \indexdef{}{command}{axiomatization}\hypertarget{command.axiomatization}{\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\
 | 
212  | 
    \indexdef{}{command}{definition}\hypertarget{command.definition}{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 28788 | 213  | 
    \indexdef{}{attribute}{defn}\hypertarget{attribute.defn}{\hyperlink{attribute.defn}{\mbox{\isa{defn}}}} & : & \isa{attribute} \\
 | 
| 40406 | 214  | 
    \indexdef{}{command}{abbreviation}\hypertarget{command.abbreviation}{\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
215  | 
    \indexdef{}{command}{print\_abbrevs}\hypertarget{command.print-abbrevs}{\hyperlink{command.print-abbrevs}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}abbrevs}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27042 | 216  | 
  \end{matharray}
 | 
217  | 
||
218  | 
These specification mechanisms provide a slightly more abstract view  | 
|
219  | 
  than the underlying primitives of \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}, \hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}} (see \secref{sec:consts}), and \hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}} (see
 | 
|
220  | 
  \secref{sec:axms-thms}).  In particular, type-inference is commonly
 | 
|
221  | 
available, and result names need not be given.  | 
|
222  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
223  | 
  \begin{railoutput}
 | 
| 42662 | 224  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
225  | 
\rail@term{\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
226  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
227  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
228  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
229  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
230  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
231  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
232  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
233  | 
\rail@nont{\isa{specs}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
234  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
235  | 
\rail@end  | 
| 42704 | 236  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
237  | 
\rail@term{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
238  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
239  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
240  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
241  | 
\rail@endbar  | 
| 42704 | 242  | 
\rail@cr{3}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
243  | 
\rail@bar  | 
| 42704 | 244  | 
\rail@nextbar{4}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
245  | 
\rail@nont{\isa{decl}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
246  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
247  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
248  | 
\rail@bar  | 
| 42704 | 249  | 
\rail@nextbar{4}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
250  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
251  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
252  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
253  | 
\rail@end  | 
| 42704 | 254  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
255  | 
\rail@term{\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
256  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
257  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
258  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
259  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
260  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
261  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
262  | 
\rail@nont{\hyperlink{syntax.mode}{\mbox{\isa{mode}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
263  | 
\rail@endbar  | 
| 42704 | 264  | 
\rail@cr{3}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
265  | 
\rail@bar  | 
| 42704 | 266  | 
\rail@nextbar{4}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
267  | 
\rail@nont{\isa{decl}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
268  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
269  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
270  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
271  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
272  | 
\rail@begin{4}{\indexdef{}{syntax}{fixes}\hypertarget{syntax.fixes}{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
273  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
274  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
275  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
276  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
277  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
278  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
279  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
280  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
281  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
282  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
283  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
284  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
285  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
286  | 
\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
287  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
288  | 
\rail@nextplus{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
289  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
290  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
291  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
292  | 
\rail@begin{3}{\isa{specs}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
293  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
294  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
295  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
296  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
297  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
298  | 
\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
299  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
300  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
301  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
302  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
303  | 
\rail@begin{2}{\isa{decl}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
304  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
305  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
306  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
307  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
308  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
309  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
310  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
311  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
312  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
313  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
314  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
315  | 
\end{railoutput}
 | 
| 27042 | 316  | 
|
317  | 
||
| 28788 | 318  | 
  \begin{description}
 | 
| 27042 | 319  | 
|
| 40406 | 320  | 
  \item \hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}
 | 
| 28788 | 321  | 
introduces several constants simultaneously and states axiomatic  | 
322  | 
properties for these. The constants are marked as being specified  | 
|
323  | 
once and for all, which prevents additional specifications being  | 
|
324  | 
issued later on.  | 
|
| 27042 | 325  | 
|
326  | 
Note that axiomatic specifications are only appropriate when  | 
|
| 28110 | 327  | 
declaring a new logical system; axiomatic specifications are  | 
328  | 
restricted to global theory contexts. Normal applications should  | 
|
329  | 
only use definitional mechanisms!  | 
|
| 27042 | 330  | 
|
| 40406 | 331  | 
  \item \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eq{\isaliteral{22}{\isachardoublequote}}} produces an
 | 
332  | 
  internal definition \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} according to the specification
 | 
|
| 27042 | 333  | 
  given as \isa{eq}, which is then turned into a proven fact.  The
 | 
334  | 
given proposition may deviate from internal meta-level equality  | 
|
335  | 
  according to the rewrite rules declared as \hyperlink{attribute.defn}{\mbox{\isa{defn}}} by the
 | 
|
| 40406 | 336  | 
  object-logic.  This usually covers object-level equality \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequote}}} and equivalence \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C65667472696768746172726F773E}{\isasymleftrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}}.  End-users normally need not
 | 
| 27042 | 337  | 
  change the \hyperlink{attribute.defn}{\mbox{\isa{defn}}} setup.
 | 
338  | 
||
339  | 
Definitions may be presented with explicit arguments on the LHS, as  | 
|
| 40406 | 340  | 
  well as additional conditions, e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}f\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} instead of
 | 
341  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ t{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}y\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isadigit{0}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ g\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ u{\isaliteral{22}{\isachardoublequote}}} instead of an
 | 
|
342  | 
  unrestricted \isa{{\isaliteral{22}{\isachardoublequote}}g\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ u{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 27042 | 343  | 
|
| 40406 | 344  | 
  \item \hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eq{\isaliteral{22}{\isachardoublequote}}} introduces a
 | 
| 28788 | 345  | 
syntactic constant which is associated with a certain term according  | 
346  | 
  to the meta-level equality \isa{eq}.
 | 
|
| 27042 | 347  | 
|
348  | 
Abbreviations participate in the usual type-inference process, but  | 
|
349  | 
are expanded before the logic ever sees them. Pretty printing of  | 
|
350  | 
terms involves higher-order rewriting with rules stemming from  | 
|
351  | 
reverted abbreviations. This needs some care to avoid overlapping  | 
|
352  | 
or looping syntactic replacements!  | 
|
353  | 
||
354  | 
  The optional \isa{mode} specification restricts output to a
 | 
|
355  | 
  particular print mode; using ``\isa{input}'' here achieves the
 | 
|
356  | 
effect of one-way abbreviations. The mode may also include an  | 
|
357  | 
  ``\hyperlink{keyword.output}{\mbox{\isa{\isakeyword{output}}}}'' qualifier that affects the concrete syntax
 | 
|
358  | 
  declared for abbreviations, cf.\ \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} in
 | 
|
359  | 
  \secref{sec:syn-trans}.
 | 
|
360  | 
||
| 40406 | 361  | 
  \item \hyperlink{command.print-abbrevs}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}abbrevs}}}} prints all constant abbreviations
 | 
| 27042 | 362  | 
of the current context.  | 
363  | 
||
| 28788 | 364  | 
  \end{description}%
 | 
| 27042 | 365  | 
\end{isamarkuptext}%
 | 
366  | 
\isamarkuptrue%  | 
|
367  | 
%  | 
|
368  | 
\isamarkupsection{Generic declarations%
 | 
|
369  | 
}  | 
|
370  | 
\isamarkuptrue%  | 
|
371  | 
%  | 
|
372  | 
\begin{isamarkuptext}%
 | 
|
373  | 
Arbitrary operations on the background context may be wrapped-up as  | 
|
374  | 
generic declaration elements. Since the underlying concept of local  | 
|
375  | 
theories may be subject to later re-interpretation, there is an  | 
|
376  | 
additional dependency on a morphism that tells the difference of the  | 
|
377  | 
original declaration context wrt.\ the application context  | 
|
378  | 
encountered later on. A fact declaration is an important special  | 
|
379  | 
case: it consists of a theorem which is applied to the context by  | 
|
380  | 
means of an attribute.  | 
|
381  | 
||
382  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 383  | 
    \indexdef{}{command}{declaration}\hypertarget{command.declaration}{\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 40784 | 384  | 
    \indexdef{}{command}{syntax\_declaration}\hypertarget{command.syntax-declaration}{\hyperlink{command.syntax-declaration}{\mbox{\isa{\isacommand{syntax{\isaliteral{5F}{\isacharunderscore}}declaration}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 40406 | 385  | 
    \indexdef{}{command}{declare}\hypertarget{command.declare}{\hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 27042 | 386  | 
  \end{matharray}
 | 
387  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
388  | 
  \begin{railoutput}
 | 
| 42704 | 389  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
390  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
391  | 
\rail@term{\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
392  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
393  | 
\rail@term{\hyperlink{command.syntax-declaration}{\mbox{\isa{\isacommand{syntax{\isaliteral{5F}{\isacharunderscore}}declaration}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
394  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
395  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
396  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
397  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
398  | 
\rail@term{\isa{\isakeyword{pervasive}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
399  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
400  | 
\rail@endbar  | 
| 42704 | 401  | 
\rail@cr{3}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
402  | 
\rail@bar  | 
| 42704 | 403  | 
\rail@nextbar{4}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
404  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
405  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
406  | 
\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
407  | 
\rail@end  | 
| 42662 | 408  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
409  | 
\rail@term{\hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
410  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
411  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
412  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
413  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
414  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
415  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
416  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
417  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
418  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
419  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
420  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
421  | 
|
| 27042 | 422  | 
|
| 28788 | 423  | 
  \begin{description}
 | 
| 27042 | 424  | 
|
| 28788 | 425  | 
  \item \hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}~\isa{d} adds the declaration
 | 
| 27042 | 426  | 
  function \isa{d} of ML type \verb|declaration|, to the current
 | 
427  | 
local theory under construction. In later application contexts, the  | 
|
428  | 
function is transformed according to the morphisms being involved in  | 
|
429  | 
the interpretation hierarchy.  | 
|
430  | 
||
| 40406 | 431  | 
  If the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}pervasive{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option is given, the corresponding
 | 
| 33516 | 432  | 
declaration is applied to all possible contexts involved, including  | 
433  | 
the global background theory.  | 
|
434  | 
||
| 40784 | 435  | 
  \item \hyperlink{command.syntax-declaration}{\mbox{\isa{\isacommand{syntax{\isaliteral{5F}{\isacharunderscore}}declaration}}}} is similar to \hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}, but is meant to affect only ``syntactic'' tools by
 | 
436  | 
convention (such as notation and type-checking information).  | 
|
437  | 
||
| 28788 | 438  | 
  \item \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}}~\isa{thms} declares theorems to the
 | 
| 27042 | 439  | 
current local theory context. No theorem binding is involved here,  | 
440  | 
  unlike \hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}} or \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}} (cf.\
 | 
|
441  | 
  \secref{sec:axms-thms}), so \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} only has the effect
 | 
|
442  | 
of applying attributes as included in the theorem specification.  | 
|
443  | 
||
| 28788 | 444  | 
  \end{description}%
 | 
| 27042 | 445  | 
\end{isamarkuptext}%
 | 
446  | 
\isamarkuptrue%  | 
|
447  | 
%  | 
|
448  | 
\isamarkupsection{Locales \label{sec:locale}%
 | 
|
449  | 
}  | 
|
450  | 
\isamarkuptrue%  | 
|
451  | 
%  | 
|
452  | 
\begin{isamarkuptext}%
 | 
|
| 33847 | 453  | 
Locales are parametric named local contexts, consisting of a list of  | 
| 27042 | 454  | 
declaration elements that are modeled after the Isar proof context  | 
455  | 
  commands (cf.\ \secref{sec:proof-context}).%
 | 
|
456  | 
\end{isamarkuptext}%
 | 
|
457  | 
\isamarkuptrue%  | 
|
458  | 
%  | 
|
| 33847 | 459  | 
\isamarkupsubsection{Locale expressions \label{sec:locale-expr}%
 | 
460  | 
}  | 
|
461  | 
\isamarkuptrue%  | 
|
462  | 
%  | 
|
463  | 
\begin{isamarkuptext}%
 | 
|
464  | 
A \emph{locale expression} denotes a structured context composed of
 | 
|
465  | 
instances of existing locales. The context consists of a list of  | 
|
466  | 
instances of declaration elements from the locales. Two locale  | 
|
467  | 
instances are equal if they are of the same locale and the  | 
|
468  | 
parameters are instantiated with equivalent terms. Declaration  | 
|
469  | 
elements from equal instances are never repeated, thus avoiding  | 
|
470  | 
duplicate declarations.  | 
|
471  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
472  | 
  \begin{railoutput}
 | 
| 42617 | 473  | 
\rail@begin{3}{\indexdef{}{syntax}{locale\_expr}\hypertarget{syntax.locale-expr}{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
474  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
475  | 
\rail@nont{\isa{instance}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
476  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
477  | 
\rail@cterm{\isa{{\isaliteral{2B}{\isacharplus}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
478  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
479  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
480  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
481  | 
\rail@term{\isa{\isakeyword{for}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
482  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
483  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
484  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
485  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
486  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
487  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
488  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
489  | 
\rail@begin{2}{\isa{instance}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
490  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
491  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
492  | 
\rail@nont{\isa{qualifier}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
493  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
494  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
495  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
496  | 
\rail@bar  | 
| 42617 | 497  | 
\rail@nont{\isa{pos{\isaliteral{5F}{\isacharunderscore}}insts}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
498  | 
\rail@nextbar{1}
 | 
| 42617 | 499  | 
\rail@nont{\isa{named{\isaliteral{5F}{\isacharunderscore}}insts}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
500  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
501  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
502  | 
\rail@begin{3}{\isa{qualifier}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
503  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
504  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
505  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
506  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
507  | 
\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
508  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
509  | 
\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
510  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
511  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
512  | 
\rail@end  | 
| 42617 | 513  | 
\rail@begin{3}{\isa{pos{\isaliteral{5F}{\isacharunderscore}}insts}}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
514  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
515  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
516  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
517  | 
\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
518  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
519  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
520  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
521  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
522  | 
\rail@end  | 
| 42617 | 523  | 
\rail@begin{2}{\isa{named{\isaliteral{5F}{\isacharunderscore}}insts}}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
524  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
525  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
526  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
527  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
528  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
529  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
530  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
531  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
532  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
533  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
534  | 
|
| 33847 | 535  | 
|
536  | 
A locale instance consists of a reference to a locale and either  | 
|
537  | 
positional or named parameter instantiations. Identical  | 
|
538  | 
instantiations (that is, those that instante a parameter by itself)  | 
|
| 40406 | 539  | 
  may be omitted.  The notation `\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}' enables to omit the
 | 
| 40256 | 540  | 
instantiation for a parameter inside a positional instantiation.  | 
| 33847 | 541  | 
|
542  | 
Terms in instantiations are from the context the locale expressions  | 
|
543  | 
is declared in. Local names may be added to this context with the  | 
|
544  | 
optional for clause. In addition, syntax declarations from one  | 
|
545  | 
instance are effective when parsing subsequent instances of the same  | 
|
546  | 
expression.  | 
|
547  | 
||
548  | 
Instances have an optional qualifier which applies to names in  | 
|
549  | 
declarations. Names include local definitions and theorem names.  | 
|
550  | 
If present, the qualifier itself is either optional  | 
|
551  | 
  (``\texttt{?}''), which means that it may be omitted on input of the
 | 
|
552  | 
  qualified name, or mandatory (``\texttt{!}'').  If neither
 | 
|
553  | 
  ``\texttt{?}'' nor ``\texttt{!}'' are present, the command's default
 | 
|
554  | 
  is used.  For \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}} and \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}
 | 
|
555  | 
  the default is ``mandatory'', for \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} and \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}} the default is ``optional''.%
 | 
|
556  | 
\end{isamarkuptext}%
 | 
|
557  | 
\isamarkuptrue%  | 
|
558  | 
%  | 
|
559  | 
\isamarkupsubsection{Locale declarations%
 | 
|
| 27042 | 560  | 
}  | 
561  | 
\isamarkuptrue%  | 
|
562  | 
%  | 
|
563  | 
\begin{isamarkuptext}%
 | 
|
564  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 565  | 
    \indexdef{}{command}{locale}\hypertarget{command.locale}{\hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
566  | 
    \indexdef{}{command}{print\_locale}\hypertarget{command.print-locale}{\hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
567  | 
    \indexdef{}{command}{print\_locales}\hypertarget{command.print-locales}{\hyperlink{command.print-locales}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locales}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
568  | 
    \indexdef{}{method}{intro\_locales}\hypertarget{method.intro-locales}{\hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}}} & : & \isa{method} \\
 | 
|
569  | 
    \indexdef{}{method}{unfold\_locales}\hypertarget{method.unfold-locales}{\hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}}} & : & \isa{method} \\
 | 
|
| 27042 | 570  | 
  \end{matharray}
 | 
571  | 
||
572  | 
  \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes}
 | 
|
| 28788 | 573  | 
  \indexisarelem{defines}\indexisarelem{notes}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
574  | 
  \begin{railoutput}
 | 
| 42662 | 575  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
576  | 
\rail@term{\hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
577  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
578  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
579  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
580  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
581  | 
\rail@nont{\hyperlink{syntax.locale}{\mbox{\isa{locale}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
582  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
583  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
584  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
585  | 
\rail@term{\isa{\isakeyword{begin}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
586  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
587  | 
\rail@end  | 
| 42662 | 588  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
589  | 
\rail@term{\hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
590  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
591  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
592  | 
\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
593  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
594  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
595  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
596  | 
\rail@begin{5}{\indexdef{}{syntax}{locale}\hypertarget{syntax.locale}{\hyperlink{syntax.locale}{\mbox{\isa{locale}}}}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
597  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
598  | 
\rail@plus  | 
| 42617 | 599  | 
\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
600  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
601  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
602  | 
\rail@nextbar{2}
 | 
| 42617 | 603  | 
\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
604  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
605  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
606  | 
\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
607  | 
\rail@plus  | 
| 42617 | 608  | 
\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
609  | 
\rail@nextplus{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
610  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
611  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
612  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
613  | 
\rail@end  | 
| 42617 | 614  | 
\rail@begin{12}{\indexdef{}{syntax}{context\_elem}\hypertarget{syntax.context-elem}{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
615  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
616  | 
\rail@term{\isa{\isakeyword{fixes}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
617  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
618  | 
\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
619  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
620  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
621  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
622  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
623  | 
\rail@term{\isa{\isakeyword{constrains}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
624  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
625  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
626  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
627  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
628  | 
\rail@nextplus{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
629  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
630  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
631  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
632  | 
\rail@term{\isa{\isakeyword{assumes}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
633  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
634  | 
\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
635  | 
\rail@nextplus{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
636  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
637  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
638  | 
\rail@nextbar{6}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
639  | 
\rail@term{\isa{\isakeyword{defines}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
640  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
641  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
642  | 
\rail@nextbar{7}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
643  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
644  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
645  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
646  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
647  | 
\rail@nextbar{7}
 | 
| 42705 | 648  | 
\rail@nont{\hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
649  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
650  | 
\rail@nextplus{8}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
651  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
652  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
653  | 
\rail@nextbar{9}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
654  | 
\rail@term{\isa{\isakeyword{notes}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
655  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
656  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
657  | 
\rail@nextbar{10}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
658  | 
\rail@nont{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
659  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
660  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
661  | 
\rail@nextplus{11}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
662  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
663  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
664  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
665  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
666  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
667  | 
|
| 27042 | 668  | 
|
| 28788 | 669  | 
  \begin{description}
 | 
| 27042 | 670  | 
|
| 40406 | 671  | 
  \item \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}loc\ {\isaliteral{3D}{\isacharequal}}\ import\ {\isaliteral{2B}{\isacharplus}}\ body{\isaliteral{22}{\isachardoublequote}}} defines a
 | 
| 27042 | 672  | 
  new locale \isa{loc} as a context consisting of a certain view of
 | 
673  | 
  existing locales (\isa{import}) plus some additional elements
 | 
|
674  | 
  (\isa{body}).  Both \isa{import} and \isa{body} are optional;
 | 
|
675  | 
  the degenerate form \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}~\isa{loc} defines an empty
 | 
|
676  | 
locale, which may still be useful to collect declarations of facts  | 
|
677  | 
later on. Type-inference on locale expressions automatically takes  | 
|
678  | 
care of the most general typing that the combined context elements  | 
|
679  | 
may acquire.  | 
|
680  | 
||
| 33847 | 681  | 
  The \isa{import} consists of a structured locale expression; see
 | 
682  | 
  \secref{sec:proof-context} above.  Its for clause defines the local
 | 
|
683  | 
  parameters of the \isa{import}.  In addition, locale parameters
 | 
|
684  | 
whose instantance is omitted automatically extend the (possibly  | 
|
685  | 
empty) for clause: they are inserted at its beginning. This means  | 
|
686  | 
that these parameters may be referred to from within the expression  | 
|
687  | 
and also in the subsequent context elements and provides a  | 
|
688  | 
notational convenience for the inheritance of parameters in locale  | 
|
689  | 
declarations.  | 
|
| 27042 | 690  | 
|
| 33847 | 691  | 
  The \isa{body} consists of context elements.
 | 
| 27042 | 692  | 
|
| 28788 | 693  | 
  \begin{description}
 | 
| 27042 | 694  | 
|
| 40406 | 695  | 
  \item \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} declares a local
 | 
696  | 
  parameter of type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} and mixfix annotation \isa{mx} (both
 | 
|
697  | 
  are optional).  The special syntax declaration ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C5354525543545552453E}{\isasymSTRUCTURE}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' means that \isa{x} may be referenced
 | 
|
| 27042 | 698  | 
implicitly in this context.  | 
699  | 
||
| 40406 | 700  | 
  \item \hyperlink{element.constrains}{\mbox{\isa{\isakeyword{constrains}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} introduces a type
 | 
701  | 
  constraint \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} on the local parameter \isa{x}.  This
 | 
|
| 38110 | 702  | 
element is deprecated. The type constraint should be introduced in  | 
| 33847 | 703  | 
  the for clause or the relevant \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} element.
 | 
| 27042 | 704  | 
|
| 40406 | 705  | 
  \item \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}
 | 
| 27042 | 706  | 
  introduces local premises, similar to \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} within a
 | 
707  | 
  proof (cf.\ \secref{sec:proof-context}).
 | 
|
708  | 
||
| 40406 | 709  | 
  \item \hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} defines a previously
 | 
| 27042 | 710  | 
  declared parameter.  This is similar to \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}} within a
 | 
711  | 
  proof (cf.\ \secref{sec:proof-context}), but \hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}}
 | 
|
712  | 
takes an equational proposition instead of variable-term pair. The  | 
|
713  | 
left-hand side of the equation may have additional arguments, e.g.\  | 
|
| 40406 | 714  | 
  ``\hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}f\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}''.
 | 
| 27042 | 715  | 
|
| 40406 | 716  | 
  \item \hyperlink{element.notes}{\mbox{\isa{\isakeyword{notes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3D}{\isacharequal}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}
 | 
| 27042 | 717  | 
reconsiders facts within a local context. Most notably, this may  | 
718  | 
include arbitrary declarations in any attribute specifications  | 
|
719  | 
  included here, e.g.\ a local \hyperlink{attribute.simp}{\mbox{\isa{simp}}} rule.
 | 
|
720  | 
||
| 28788 | 721  | 
  The initial \isa{import} specification of a locale expression
 | 
722  | 
maintains a dynamic relation to the locales being referenced  | 
|
723  | 
(benefiting from any later fact declarations in the obvious manner).  | 
|
| 27042 | 724  | 
|
| 28788 | 725  | 
  \end{description}
 | 
| 27042 | 726  | 
|
| 40406 | 727  | 
  Note that ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C49533E}{\isasymIS}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' patterns given
 | 
| 27042 | 728  | 
  in the syntax of \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} and \hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}} above
 | 
729  | 
are illegal in locale definitions. In the long goal format of  | 
|
730  | 
  \secref{sec:goals}, term bindings may be included as expected,
 | 
|
731  | 
though.  | 
|
732  | 
||
| 33847 | 733  | 
\medskip Locale specifications are ``closed up'' by  | 
| 40406 | 734  | 
  turning the given text into a predicate definition \isa{loc{\isaliteral{5F}{\isacharunderscore}}axioms} and deriving the original assumptions as local lemmas
 | 
| 27042 | 735  | 
(modulo local definitions). The predicate statement covers only the  | 
736  | 
newly specified assumptions, omitting the content of included locale  | 
|
737  | 
expressions. The full cumulative view is only provided on export,  | 
|
738  | 
  involving another predicate \isa{loc} that refers to the complete
 | 
|
739  | 
specification text.  | 
|
740  | 
||
741  | 
In any case, the predicate arguments are those locale parameters  | 
|
742  | 
that actually occur in the respective piece of text. Also note that  | 
|
743  | 
these predicates operate at the meta-level in theory, but the locale  | 
|
744  | 
packages attempts to internalize statements according to the  | 
|
| 40406 | 745  | 
  object-logic setup (e.g.\ replacing \isa{{\isaliteral{5C3C416E643E}{\isasymAnd}}} by \isa{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}}, and
 | 
746  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} in HOL; see also
 | 
|
747  | 
  \secref{sec:object-logic}).  Separate introduction rules \isa{loc{\isaliteral{5F}{\isacharunderscore}}axioms{\isaliteral{2E}{\isachardot}}intro} and \isa{loc{\isaliteral{2E}{\isachardot}}intro} are provided as well.
 | 
|
| 27042 | 748  | 
|
| 40406 | 749  | 
  \item \hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}locale{\isaliteral{22}{\isachardoublequote}}} prints the
 | 
| 33868 | 750  | 
  contents of the named locale.  The command omits \hyperlink{element.notes}{\mbox{\isa{\isakeyword{notes}}}}
 | 
| 40406 | 751  | 
  elements by default.  Use \hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} to
 | 
| 33868 | 752  | 
have them included.  | 
| 27042 | 753  | 
|
| 40406 | 754  | 
  \item \hyperlink{command.print-locales}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locales}}}} prints the names of all locales
 | 
| 27042 | 755  | 
of the current theory.  | 
756  | 
||
| 40406 | 757  | 
  \item \hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}} and \hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}}
 | 
| 27042 | 758  | 
repeatedly expand all introduction rules of locale predicates of the  | 
| 40406 | 759  | 
  theory.  While \hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}} only applies the \isa{loc{\isaliteral{2E}{\isachardot}}intro} introduction rules and therefore does not decend to
 | 
760  | 
  assumptions, \hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}} is more aggressive and applies
 | 
|
761  | 
  \isa{loc{\isaliteral{5F}{\isacharunderscore}}axioms{\isaliteral{2E}{\isachardot}}intro} as well.  Both methods are aware of locale
 | 
|
| 28728 | 762  | 
specifications entailed by the context, both from target statements,  | 
763  | 
and from interpretations (see below). New goals that are entailed  | 
|
764  | 
by the current context are discharged automatically.  | 
|
| 27042 | 765  | 
|
| 28788 | 766  | 
  \end{description}%
 | 
| 27042 | 767  | 
\end{isamarkuptext}%
 | 
768  | 
\isamarkuptrue%  | 
|
769  | 
%  | 
|
| 33847 | 770  | 
\isamarkupsubsection{Locale interpretations%
 | 
| 27042 | 771  | 
}  | 
772  | 
\isamarkuptrue%  | 
|
773  | 
%  | 
|
774  | 
\begin{isamarkuptext}%
 | 
|
| 33847 | 775  | 
Locale expressions may be instantiated, and the instantiated facts  | 
776  | 
added to the current context. This requires a proof of the  | 
|
777  | 
  instantiated specification and is called \emph{locale
 | 
|
778  | 
  interpretation}.  Interpretation is possible in locales \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}, theories (command \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}) and
 | 
|
779  | 
  also within a proof body (command \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}).
 | 
|
| 27042 | 780  | 
|
781  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 782  | 
    \indexdef{}{command}{interpretation}\hypertarget{command.interpretation}{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
783  | 
    \indexdef{}{command}{interpret}\hypertarget{command.interpret}{\hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 
41434
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
784  | 
    \indexdef{}{command}{sublocale}\hypertarget{command.sublocale}{\hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 41435 | 785  | 
    \indexdef{}{command}{print\_dependencies}\hypertarget{command.print-dependencies}{\hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 40406 | 786  | 
    \indexdef{}{command}{print\_interps}\hypertarget{command.print-interps}{\hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}interps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 27042 | 787  | 
  \end{matharray}
 | 
788  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
789  | 
  \begin{railoutput}
 | 
| 42662 | 790  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
791  | 
\rail@term{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}}[]
 | 
| 42617 | 792  | 
\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
793  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
794  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
795  | 
\rail@nont{\isa{equations}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
796  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
797  | 
\rail@end  | 
| 42662 | 798  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
799  | 
\rail@term{\hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}}[]
 | 
| 42617 | 800  | 
\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
801  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
802  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
803  | 
\rail@nont{\isa{equations}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
804  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
805  | 
\rail@end  | 
| 42704 | 806  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
807  | 
\rail@term{\hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
808  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
809  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
810  | 
\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
811  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
812  | 
\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
813  | 
\rail@endbar  | 
| 42617 | 814  | 
\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[]
 | 
| 42704 | 815  | 
\rail@cr{3}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
816  | 
\rail@bar  | 
| 42704 | 817  | 
\rail@nextbar{4}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
818  | 
\rail@nont{\isa{equations}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
819  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
820  | 
\rail@end  | 
| 42662 | 821  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
822  | 
\rail@term{\hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
823  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
824  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
825  | 
\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
826  | 
\rail@endbar  | 
| 42617 | 827  | 
\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
828  | 
\rail@end  | 
| 42662 | 829  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
830  | 
\rail@term{\hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}interps}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
831  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
832  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
833  | 
\rail@begin{3}{\isa{equations}}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
834  | 
\rail@term{\isa{\isakeyword{where}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
835  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
836  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
837  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
838  | 
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
839  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
840  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
841  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
842  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
843  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
844  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
845  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
846  | 
|
| 27042 | 847  | 
|
| 28788 | 848  | 
  \begin{description}
 | 
| 27042 | 849  | 
|
| 40406 | 850  | 
  \item \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}expr\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eqns{\isaliteral{22}{\isachardoublequote}}}
 | 
| 33847 | 851  | 
  interprets \isa{expr} in the theory.  The command generates proof
 | 
852  | 
obligations for the instantiated specifications (assumes and defines  | 
|
853  | 
elements). Once these are discharged by the user, instantiated  | 
|
854  | 
facts are added to the theory in a post-processing phase.  | 
|
855  | 
||
856  | 
Additional equations, which are unfolded during  | 
|
| 27042 | 857  | 
  post-processing, may be given after the keyword \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}}.
 | 
858  | 
This is useful for interpreting concepts introduced through  | 
|
| 
41434
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
859  | 
definitions. The equations must be proved.  | 
| 27042 | 860  | 
|
861  | 
The command is aware of interpretations already active in the  | 
|
| 
28085
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
27834 
diff
changeset
 | 
862  | 
theory, but does not simplify the goal automatically. In order to  | 
| 40406 | 863  | 
  simplify the proof obligations use methods \hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}}
 | 
864  | 
  or \hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}}.  Post-processing is not applied to
 | 
|
| 
28085
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
27834 
diff
changeset
 | 
865  | 
facts of interpretations that are already active. This avoids  | 
| 
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
27834 
diff
changeset
 | 
866  | 
duplication of interpreted facts, in particular. Note that, in the  | 
| 
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
27834 
diff
changeset
 | 
867  | 
case of a locale with import, parts of the interpretation may  | 
| 
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
27834 
diff
changeset
 | 
868  | 
already be active. The command will only process facts for new  | 
| 
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
27834 
diff
changeset
 | 
869  | 
parts.  | 
| 27042 | 870  | 
|
871  | 
Adding facts to locales has the effect of adding interpreted facts  | 
|
| 
41434
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
872  | 
to the theory for all interpretations as well. That is,  | 
| 27042 | 873  | 
interpretations dynamically participate in any facts added to  | 
| 
41434
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
874  | 
locales. Note that if a theory inherits additional facts for a  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
875  | 
locale through one parent and an interpretation of that locale  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
876  | 
through another parent, the additional facts will not be  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
877  | 
interpreted.  | 
| 27042 | 878  | 
|
| 40406 | 879  | 
  \item \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}expr\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eqns{\isaliteral{22}{\isachardoublequote}}} interprets
 | 
| 38110 | 880  | 
  \isa{expr} in the proof context and is otherwise similar to
 | 
881  | 
interpretation in theories. Note that rewrite rules given to  | 
|
| 
41434
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
882  | 
  \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}} after the \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} keyword should be
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
883  | 
explicitly universally quantified.  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
884  | 
|
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
885  | 
  \item \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ expr\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eqns{\isaliteral{22}{\isachardoublequote}}}
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
886  | 
  interprets \isa{expr} in the locale \isa{name}.  A proof that
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
887  | 
  the specification of \isa{name} implies the specification of
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
888  | 
  \isa{expr} is required.  As in the localized version of the
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
889  | 
  theorem command, the proof is in the context of \isa{name}.  After
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
890  | 
  the proof obligation has been discharged, the facts of \isa{expr}
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
891  | 
  become part of locale \isa{name} as \emph{derived} context
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
892  | 
  elements and are available when the context \isa{name} is
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
893  | 
subsequently entered. Note that, like import, this is dynamic:  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
894  | 
  facts added to a locale part of \isa{expr} after interpretation
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
895  | 
  become also available in \isa{name}.
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
896  | 
|
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
897  | 
  Only specification fragments of \isa{expr} that are not already
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
898  | 
  part of \isa{name} (be it imported, derived or a derived fragment
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
899  | 
of the import) are considered in this process. This enables  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
900  | 
circular interpretations provided that no infinite chains are  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
901  | 
generated in the locale hierarchy.  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
902  | 
|
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
903  | 
  If interpretations of \isa{name} exist in the current theory, the
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
904  | 
  command adds interpretations for \isa{expr} as well, with the same
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
905  | 
  qualifier, although only for fragments of \isa{expr} that are not
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
906  | 
interpreted in the theory already.  | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
907  | 
|
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
908  | 
  Equations given after \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} amend the morphism through
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
909  | 
  which \isa{expr} is interpreted.  This enables to map definitions
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
910  | 
  from the interpreted locales to entities of \isa{name}.  This
 | 
| 
 
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
 
ballarin 
parents: 
41249 
diff
changeset
 | 
911  | 
feature is experimental.  | 
| 27042 | 912  | 
|
| 41435 | 913  | 
  \item \hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}expr{\isaliteral{22}{\isachardoublequote}}} is useful for
 | 
914  | 
  understanding the effect of an interpretation of \isa{{\isaliteral{22}{\isachardoublequote}}expr{\isaliteral{22}{\isachardoublequote}}}.  It
 | 
|
915  | 
lists all locale instances for which interpretations would be added  | 
|
916  | 
  to the current context.  Variant \hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} prints all locale instances that
 | 
|
917  | 
would be considered for interpretation, and would be interpreted in  | 
|
918  | 
an empty context (that is, without interpretations).  | 
|
919  | 
||
| 40406 | 920  | 
  \item \hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}interps}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}locale{\isaliteral{22}{\isachardoublequote}}} lists all
 | 
921  | 
  interpretations of \isa{{\isaliteral{22}{\isachardoublequote}}locale{\isaliteral{22}{\isachardoublequote}}} in the current theory or proof
 | 
|
| 38110 | 922  | 
  context, including those due to a combination of a \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}} or \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}} and one or several
 | 
923  | 
  \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}} declarations.
 | 
|
| 33868 | 924  | 
|
| 28788 | 925  | 
  \end{description}
 | 
| 27042 | 926  | 
|
927  | 
  \begin{warn}
 | 
|
928  | 
Since attributes are applied to interpreted theorems,  | 
|
929  | 
interpretation may modify the context of common proof tools, e.g.\  | 
|
| 33868 | 930  | 
the Simplifier or Classical Reasoner. As the behavior of such  | 
931  | 
    tools is \emph{not} stable under interpretation morphisms, manual
 | 
|
932  | 
declarations might have to be added to the target context of the  | 
|
933  | 
interpretation to revert such declarations.  | 
|
| 27042 | 934  | 
  \end{warn}
 | 
935  | 
||
936  | 
  \begin{warn}
 | 
|
| 38110 | 937  | 
An interpretation in a theory or proof context may subsume previous  | 
| 27042 | 938  | 
interpretations. This happens if the same specification fragment  | 
939  | 
is interpreted twice and the instantiation of the second  | 
|
940  | 
interpretation is more general than the interpretation of the  | 
|
| 33847 | 941  | 
first. The locale package does not attempt to remove subsumed  | 
942  | 
interpretations.  | 
|
| 27042 | 943  | 
  \end{warn}%
 | 
944  | 
\end{isamarkuptext}%
 | 
|
945  | 
\isamarkuptrue%  | 
|
946  | 
%  | 
|
947  | 
\isamarkupsection{Classes \label{sec:class}%
 | 
|
948  | 
}  | 
|
949  | 
\isamarkuptrue%  | 
|
950  | 
%  | 
|
951  | 
\begin{isamarkuptext}%
 | 
|
952  | 
A class is a particular locale with \emph{exactly one} type variable
 | 
|
| 40406 | 953  | 
  \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}}.  Beyond the underlying locale, a corresponding type class
 | 
| 27042 | 954  | 
is established which is interpreted logically as axiomatic type  | 
955  | 
  class \cite{Wenzel:1997:TPHOL} whose logical content are the
 | 
|
956  | 
assumptions of the locale. Thus, classes provide the full  | 
|
957  | 
generality of locales combined with the commodity of type classes  | 
|
958  | 
  (notably type-inference).  See \cite{isabelle-classes} for a short
 | 
|
959  | 
tutorial.  | 
|
960  | 
||
961  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 962  | 
    \indexdef{}{command}{class}\hypertarget{command.class}{\hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
963  | 
    \indexdef{}{command}{instantiation}\hypertarget{command.instantiation}{\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
964  | 
    \indexdef{}{command}{instance}\hypertarget{command.instance}{\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 42626 | 965  | 
    \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 40406 | 966  | 
    \indexdef{}{command}{subclass}\hypertarget{command.subclass}{\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
967  | 
    \indexdef{}{command}{print\_classes}\hypertarget{command.print-classes}{\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}classes}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
968  | 
    \indexdef{}{command}{class\_deps}\hypertarget{command.class-deps}{\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isaliteral{5F}{\isacharunderscore}}deps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
969  | 
    \indexdef{}{method}{intro\_classes}\hypertarget{method.intro-classes}{\hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}}} & : & \isa{method} \\
 | 
|
| 27042 | 970  | 
  \end{matharray}
 | 
971  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
972  | 
  \begin{railoutput}
 | 
| 42704 | 973  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
974  | 
\rail@term{\hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}}[]
 | 
| 42704 | 975  | 
\rail@nont{\isa{class{\isaliteral{5F}{\isacharunderscore}}spec}}[]
 | 
976  | 
\rail@bar  | 
|
977  | 
\rail@nextbar{1}
 | 
|
978  | 
\rail@term{\isa{\isakeyword{begin}}}[]
 | 
|
979  | 
\rail@endbar  | 
|
980  | 
\rail@end  | 
|
981  | 
\rail@begin{5}{\isa{class{\isaliteral{5F}{\isacharunderscore}}spec}}
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
982  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
983  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
984  | 
\rail@bar  | 
| 42617 | 985  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
986  | 
\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
987  | 
\rail@plus  | 
| 42617 | 988  | 
\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
989  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
990  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
991  | 
\rail@nextbar{2}
 | 
| 42617 | 992  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
993  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
994  | 
\rail@plus  | 
| 42617 | 995  | 
\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
996  | 
\rail@nextplus{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
997  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
998  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
999  | 
\rail@end  | 
| 42662 | 1000  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1001  | 
\rail@term{\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1002  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1003  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1004  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1005  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1006  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1007  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1008  | 
\rail@nont{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1009  | 
\rail@term{\isa{\isakeyword{begin}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1010  | 
\rail@end  | 
| 42704 | 1011  | 
\rail@begin{5}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1012  | 
\rail@term{\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}}[]
 | 
| 42704 | 1013  | 
\rail@bar  | 
1014  | 
\rail@nextbar{1}
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1015  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1016  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 42704 | 1017  | 
\rail@nextplus{2}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1018  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1019  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1020  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1021  | 
\rail@nont{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}[]
 | 
| 42704 | 1022  | 
\rail@nextbar{3}
 | 
1023  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
|
1024  | 
\rail@bar  | 
|
1025  | 
\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[]
 | 
|
1026  | 
\rail@nextbar{4}
 | 
|
1027  | 
\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[]
 | 
|
1028  | 
\rail@endbar  | 
|
1029  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
|
1030  | 
\rail@endbar  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1031  | 
\rail@end  | 
| 42662 | 1032  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1033  | 
\rail@term{\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1034  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1035  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1036  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1037  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1038  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1039  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1040  | 
\end{railoutput}
 | 
| 42617 | 1041  | 
|
| 27042 | 1042  | 
|
| 28788 | 1043  | 
  \begin{description}
 | 
| 27042 | 1044  | 
|
| 40406 | 1045  | 
  \item \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3D}{\isacharequal}}\ superclasses\ {\isaliteral{2B}{\isacharplus}}\ body{\isaliteral{22}{\isachardoublequote}}} defines
 | 
| 27042 | 1046  | 
  a new class \isa{c}, inheriting from \isa{superclasses}.  This
 | 
1047  | 
  introduces a locale \isa{c} with import of all locales \isa{superclasses}.
 | 
|
1048  | 
||
1049  | 
  Any \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} in \isa{body} are lifted to the global
 | 
|
| 40406 | 1050  | 
  theory level (\emph{class operations} \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ f\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} of class \isa{c}), mapping the local type parameter
 | 
1051  | 
  \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}} to a schematic type variable \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 27042 | 1052  | 
|
1053  | 
  Likewise, \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} in \isa{body} are also lifted,
 | 
|
| 40406 | 1054  | 
  mapping each local parameter \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} to its
 | 
1055  | 
  corresponding global constant \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}.  The
 | 
|
1056  | 
  corresponding introduction rule is provided as \isa{c{\isaliteral{5F}{\isacharunderscore}}class{\isaliteral{5F}{\isacharunderscore}}axioms{\isaliteral{2E}{\isachardot}}intro}.  This rule should be rarely needed directly
 | 
|
1057  | 
  --- the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}} method takes care of the details of
 | 
|
| 27042 | 1058  | 
class membership proofs.  | 
1059  | 
||
| 40406 | 1060  | 
  \item \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}s\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} opens a theory target (cf.\ \secref{sec:target}) which
 | 
1061  | 
  allows to specify class operations \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ f\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} corresponding
 | 
|
1062  | 
  to sort \isa{s} at the particular type instance \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}}.  A plain \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} command in the
 | 
|
| 28788 | 1063  | 
target body poses a goal stating these type arities. The target is  | 
1064  | 
  concluded by an \indexref{local}{command}{end}\hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}} command.
 | 
|
| 27042 | 1065  | 
|
1066  | 
Note that a list of simultaneous type constructors may be given;  | 
|
| 31914 | 1067  | 
this corresponds nicely to mutually recursive type definitions, e.g.\  | 
| 27042 | 1068  | 
in Isabelle/HOL.  | 
1069  | 
||
| 28788 | 1070  | 
  \item \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} in an instantiation target body sets
 | 
| 40406 | 1071  | 
  up a goal stating the type arities claimed at the opening \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}.  The proof would usually proceed by \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}}, and then establish the characteristic theorems of
 | 
| 27042 | 1072  | 
the type classes involved. After finishing the proof, the  | 
1073  | 
background theory will be augmented by the proven type arities.  | 
|
1074  | 
||
| 40406 | 1075  | 
  On the theory level, \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}s{\isaliteral{22}{\isachardoublequote}}} provides a convenient way to instantiate a type class with no
 | 
| 37115 | 1076  | 
need to specify operations: one can continue with the  | 
| 31681 | 1077  | 
instantiation proof immediately.  | 
1078  | 
||
| 28788 | 1079  | 
  \item \hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}}~\isa{c} in a class context for class
 | 
| 27042 | 1080  | 
  \isa{d} sets up a goal stating that class \isa{c} is logically
 | 
1081  | 
  contained in class \isa{d}.  After finishing the proof, class
 | 
|
1082  | 
  \isa{d} is proven to be subclass \isa{c} and the locale \isa{c} is interpreted into \isa{d} simultaneously.
 | 
|
1083  | 
||
| 31681 | 1084  | 
A weakend form of this is available through a further variant of  | 
| 40406 | 1085  | 
  \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}:  \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} opens
 | 
1086  | 
  a proof that class \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} implies \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} without reference
 | 
|
| 31681 | 1087  | 
to the underlying locales; this is useful if the properties to prove  | 
1088  | 
the logical connection are not sufficent on the locale level but on  | 
|
1089  | 
the theory level.  | 
|
1090  | 
||
| 40406 | 1091  | 
  \item \hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}classes}}}} prints all classes in the current
 | 
| 27042 | 1092  | 
theory.  | 
1093  | 
||
| 40406 | 1094  | 
  \item \hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isaliteral{5F}{\isacharunderscore}}deps}}}} visualizes all classes and their
 | 
| 29706 | 1095  | 
subclass relations as a Hasse diagram.  | 
1096  | 
||
| 40406 | 1097  | 
  \item \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}} repeatedly expands all class
 | 
| 27042 | 1098  | 
introduction rules of this theory. Note that this method usually  | 
1099  | 
needs not be named explicitly, as it is already included in the  | 
|
1100  | 
  default proof step (e.g.\ of \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}).  In particular,
 | 
|
1101  | 
instantiation of trivial (syntactic) classes may be performed by a  | 
|
| 40406 | 1102  | 
  single ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' proof step.
 | 
| 26870 | 1103  | 
|
| 28788 | 1104  | 
  \end{description}%
 | 
| 26870 | 1105  | 
\end{isamarkuptext}%
 | 
1106  | 
\isamarkuptrue%  | 
|
1107  | 
%  | 
|
| 27042 | 1108  | 
\isamarkupsubsection{The class target%
 | 
1109  | 
}  | 
|
1110  | 
\isamarkuptrue%  | 
|
1111  | 
%  | 
|
1112  | 
\begin{isamarkuptext}%
 | 
|
1113  | 
%FIXME check  | 
|
1114  | 
||
1115  | 
  A named context may refer to a locale (cf.\ \secref{sec:target}).
 | 
|
1116  | 
  If this locale is also a class \isa{c}, apart from the common
 | 
|
1117  | 
locale target behaviour the following happens.  | 
|
1118  | 
||
1119  | 
  \begin{itemize}
 | 
|
1120  | 
||
| 40406 | 1121  | 
  \item Local constant declarations \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} referring to the
 | 
1122  | 
  local type parameter \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}} and local parameters \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1123  | 
  are accompanied by theory-level constants \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1124  | 
  referring to theory-level class operations \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 27042 | 1125  | 
|
1126  | 
\item Local theorem bindings are lifted as are assumptions.  | 
|
1127  | 
||
| 40406 | 1128  | 
  \item Local syntax refers to local operations \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} and
 | 
1129  | 
  global operations \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} uniformly.  Type inference
 | 
|
| 27042 | 1130  | 
resolves ambiguities. In rare cases, manual type annotations are  | 
1131  | 
needed.  | 
|
1132  | 
||
1133  | 
  \end{itemize}%
 | 
|
1134  | 
\end{isamarkuptext}%
 | 
|
1135  | 
\isamarkuptrue%  | 
|
1136  | 
%  | 
|
| 
37768
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1137  | 
\isamarkupsubsection{Co-regularity of type classes and arities%
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1138  | 
}  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1139  | 
\isamarkuptrue%  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1140  | 
%  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1141  | 
\begin{isamarkuptext}%
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1142  | 
The class relation together with the collection of  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1143  | 
type-constructor arities must obey the principle of  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1144  | 
  \emph{co-regularity} as defined below.
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1145  | 
|
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1146  | 
\medskip For the subsequent formulation of co-regularity we assume  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1147  | 
that the class relation is closed by transitivity and reflexivity.  | 
| 40406 | 1148  | 
  Moreover the collection of arities \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s{\isaliteral{29}{\isacharparenright}}c{\isaliteral{22}{\isachardoublequote}}} is
 | 
1149  | 
  completed such that \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s{\isaliteral{29}{\isacharparenright}}c{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}}
 | 
|
1150  | 
  implies \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s{\isaliteral{29}{\isacharparenright}}c{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} for all such declarations.
 | 
|
| 
37768
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1151  | 
|
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1152  | 
Treating sorts as finite sets of classes (meaning the intersection),  | 
| 40406 | 1153  | 
  the class relation \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} is extended to sorts as
 | 
| 
37768
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1154  | 
follows:  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1155  | 
\[  | 
| 40406 | 1156  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 
37768
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1157  | 
\]  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1158  | 
|
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1159  | 
This relation on sorts is further extended to tuples of sorts (of  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1160  | 
the same length) in the component-wise way.  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1161  | 
|
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1162  | 
\smallskip Co-regularity of the class relation together with the  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1163  | 
arities relation means:  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1164  | 
\[  | 
| 40406 | 1165  | 
    \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 
37768
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1166  | 
\]  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1167  | 
\noindent for all such arities. In other words, whenever the result  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1168  | 
classes of some type-constructor arities are related, then the  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1169  | 
argument sorts need to be related in the same way.  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1170  | 
|
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1171  | 
\medskip Co-regularity is a very fundamental property of the  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1172  | 
order-sorted algebra of types. For example, it entails principle  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1173  | 
  types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}.%
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1174  | 
\end{isamarkuptext}%
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1175  | 
\isamarkuptrue%  | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1176  | 
%  | 
| 27042 | 1177  | 
\isamarkupsection{Unrestricted overloading%
 | 
1178  | 
}  | 
|
1179  | 
\isamarkuptrue%  | 
|
1180  | 
%  | 
|
1181  | 
\begin{isamarkuptext}%
 | 
|
1182  | 
Isabelle/Pure's definitional schemes support certain forms of  | 
|
| 31047 | 1183  | 
  overloading (see \secref{sec:consts}).  Overloading means that a
 | 
| 40406 | 1184  | 
  constant being declared as \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ decl{\isaliteral{22}{\isachardoublequote}}} may be
 | 
| 31047 | 1185  | 
defined separately on type instances  | 
| 40406 | 1186  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ decl{\isaliteral{22}{\isachardoublequote}}}
 | 
| 31047 | 1187  | 
  for each type constructor \isa{t}.  At most occassions
 | 
| 27042 | 1188  | 
overloading will be used in a Haskell-like fashion together with  | 
1189  | 
  type classes by means of \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} (see
 | 
|
1190  | 
  \secref{sec:class}).  Sometimes low-level overloading is desirable.
 | 
|
1191  | 
  The \hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}} target provides a convenient view for
 | 
|
1192  | 
end-users.  | 
|
1193  | 
||
1194  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 1195  | 
    \indexdef{}{command}{overloading}\hypertarget{command.overloading}{\hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 27042 | 1196  | 
  \end{matharray}
 | 
1197  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1198  | 
  \begin{railoutput}
 | 
| 42704 | 1199  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1200  | 
\rail@term{\hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1201  | 
\rail@plus  | 
| 42704 | 1202  | 
\rail@nont{\isa{spec}}[]
 | 
1203  | 
\rail@nextplus{1}
 | 
|
1204  | 
\rail@endplus  | 
|
1205  | 
\rail@term{\isa{\isakeyword{begin}}}[]
 | 
|
1206  | 
\rail@end  | 
|
1207  | 
\rail@begin{2}{\isa{spec}}
 | 
|
| 42617 | 1208  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1209  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1210  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 42704 | 1211  | 
\rail@nextbar{1}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1212  | 
\rail@term{\isa{{\isaliteral{5C3C65717569763E}{\isasymequiv}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1213  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1214  | 
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1215  | 
\rail@bar  | 
| 42704 | 1216  | 
\rail@nextbar{1}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1217  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1218  | 
\rail@term{\isa{\isakeyword{unchecked}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1219  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1220  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1221  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1222  | 
\end{railoutput}
 | 
| 42617 | 1223  | 
|
| 27042 | 1224  | 
|
| 28788 | 1225  | 
  \begin{description}
 | 
| 27042 | 1226  | 
|
| 40406 | 1227  | 
  \item \hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 27042 | 1228  | 
  opens a theory target (cf.\ \secref{sec:target}) which allows to
 | 
1229  | 
specify constants with overloaded definitions. These are identified  | 
|
| 40406 | 1230  | 
  by an explicitly given mapping from variable names \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} to
 | 
1231  | 
  constants \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} at particular type instances.  The
 | 
|
| 28788 | 1232  | 
definitions themselves are established using common specification  | 
| 40406 | 1233  | 
  tools, using the names \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} as reference to the
 | 
| 28788 | 1234  | 
  corresponding constants.  The target is concluded by \hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}}.
 | 
| 27042 | 1235  | 
|
| 40406 | 1236  | 
  A \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}unchecked{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option disables global dependency checks for
 | 
| 27042 | 1237  | 
the corresponding definition, which is occasionally useful for  | 
| 31047 | 1238  | 
  exotic overloading (see \secref{sec:consts} for a precise description).
 | 
1239  | 
It is at the discretion of the user to avoid  | 
|
| 27042 | 1240  | 
malformed theory specifications!  | 
1241  | 
||
| 28788 | 1242  | 
  \end{description}%
 | 
| 27042 | 1243  | 
\end{isamarkuptext}%
 | 
1244  | 
\isamarkuptrue%  | 
|
1245  | 
%  | 
|
1246  | 
\isamarkupsection{Incorporating ML code \label{sec:ML}%
 | 
|
1247  | 
}  | 
|
1248  | 
\isamarkuptrue%  | 
|
1249  | 
%  | 
|
1250  | 
\begin{isamarkuptext}%
 | 
|
1251  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 1252  | 
    \indexdef{}{command}{use}\hypertarget{command.use}{\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
1253  | 
    \indexdef{}{command}{ML}\hypertarget{command.ML}{\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1254  | 
    \indexdef{}{command}{ML\_prf}\hypertarget{command.ML-prf}{\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1255  | 
    \indexdef{}{command}{ML\_val}\hypertarget{command.ML-val}{\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1256  | 
    \indexdef{}{command}{ML\_command}\hypertarget{command.ML-command}{\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1257  | 
    \indexdef{}{command}{setup}\hypertarget{command.setup}{\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1258  | 
    \indexdef{}{command}{local\_setup}\hypertarget{command.local-setup}{\hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1259  | 
    \indexdef{}{command}{attribute\_setup}\hypertarget{command.attribute-setup}{\hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27042 | 1260  | 
  \end{matharray}
 | 
1261  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1262  | 
  \begin{railoutput}
 | 
| 42662 | 1263  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1264  | 
\rail@term{\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1265  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1266  | 
\rail@end  | 
| 42662 | 1267  | 
\rail@begin{6}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1268  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1269  | 
\rail@term{\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1270  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1271  | 
\rail@term{\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1272  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1273  | 
\rail@term{\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1274  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1275  | 
\rail@term{\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1276  | 
\rail@nextbar{4}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1277  | 
\rail@term{\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1278  | 
\rail@nextbar{5}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1279  | 
\rail@term{\hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isaliteral{5F}{\isacharunderscore}}setup}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1280  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1281  | 
\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1282  | 
\rail@end  | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
1283  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1284  | 
\rail@term{\hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1285  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1286  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1287  | 
\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[]
 | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
1288  | 
\rail@bar  | 
| 
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
1289  | 
\rail@nextbar{1}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1290  | 
\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[]
 | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42705 
diff
changeset
 | 
1291  | 
\rail@endbar  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1292  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1293  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1294  | 
|
| 27042 | 1295  | 
|
| 28788 | 1296  | 
  \begin{description}
 | 
| 27042 | 1297  | 
|
| 40406 | 1298  | 
  \item \hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}file{\isaliteral{22}{\isachardoublequote}}} reads and executes ML
 | 
1299  | 
  commands from \isa{{\isaliteral{22}{\isachardoublequote}}file{\isaliteral{22}{\isachardoublequote}}}.  The current theory context is passed
 | 
|
| 28788 | 1300  | 
down to the ML toplevel and may be modified, using \verb|Context.>>| or derived ML commands. The file name is checked with  | 
| 27042 | 1301  | 
  the \indexref{}{keyword}{uses}\hyperlink{keyword.uses}{\mbox{\isa{\isakeyword{uses}}}} dependency declaration given in the theory
 | 
1302  | 
  header (see also \secref{sec:begin-thy}).
 | 
|
| 28281 | 1303  | 
|
1304  | 
Top-level ML bindings are stored within the (global or local) theory  | 
|
1305  | 
context.  | 
|
| 27042 | 1306  | 
|
| 40406 | 1307  | 
  \item \hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} is similar to \hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}},
 | 
1308  | 
  but executes ML commands directly from the given \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 28788 | 1309  | 
Top-level ML bindings are stored within the (global or local) theory  | 
1310  | 
context.  | 
|
| 28281 | 1311  | 
|
| 40406 | 1312  | 
  \item \hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}} is analogous to \hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}} but works
 | 
| 28788 | 1313  | 
within a proof context.  | 
| 28281 | 1314  | 
|
1315  | 
Top-level ML bindings are stored within the proof context in a  | 
|
1316  | 
purely sequential fashion, disregarding the nested proof structure.  | 
|
| 40406 | 1317  | 
  ML bindings introduced by \hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}} are discarded at the
 | 
| 28281 | 1318  | 
end of the proof.  | 
| 27042 | 1319  | 
|
| 40406 | 1320  | 
  \item \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}} and \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}} are diagnostic
 | 
| 28788 | 1321  | 
  versions of \hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}, which means that the context may not be
 | 
| 40406 | 1322  | 
  updated.  \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}} echos the bindings produced at the ML
 | 
1323  | 
  toplevel, but \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}} is silent.
 | 
|
| 27042 | 1324  | 
|
| 40406 | 1325  | 
  \item \hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} changes the current theory
 | 
1326  | 
  context by applying \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}}, which refers to an ML expression
 | 
|
| 30463 | 1327  | 
of type \verb|theory -> theory|. This enables to initialize  | 
1328  | 
any object-logic specific tools and packages written in ML, for  | 
|
1329  | 
example.  | 
|
1330  | 
||
| 40406 | 1331  | 
  \item \hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isaliteral{5F}{\isacharunderscore}}setup}}}} is similar to \hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}} for
 | 
| 30463 | 1332  | 
a local theory context, and an ML expression of type \verb|local_theory -> local_theory|. This allows to  | 
1333  | 
invoke local theory specification packages without going through  | 
|
1334  | 
concrete outer syntax, for example.  | 
|
| 27042 | 1335  | 
|
| 40406 | 1336  | 
  \item \hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ text\ description{\isaliteral{22}{\isachardoublequote}}}
 | 
1337  | 
  defines an attribute in the current theory.  The given \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} has to be an ML expression of type
 | 
|
| 30527 | 1338  | 
\verb|attribute context_parser|, cf.\ basic parsers defined in  | 
1339  | 
structure \verb|Args| and \verb|Attrib|.  | 
|
1340  | 
||
1341  | 
In principle, attributes can operate both on a given theorem and the  | 
|
1342  | 
implicit context, although in practice only one is modified and the  | 
|
1343  | 
other serves as parameter. Here are examples for these two cases:  | 
|
1344  | 
||
1345  | 
  \end{description}%
 | 
|
1346  | 
\end{isamarkuptext}%
 | 
|
1347  | 
\isamarkuptrue%  | 
|
1348  | 
%  | 
|
1349  | 
\isadelimML  | 
|
| 42704 | 1350  | 
\ \ %  | 
| 30527 | 1351  | 
\endisadelimML  | 
1352  | 
%  | 
|
1353  | 
\isatagML  | 
|
| 40406 | 1354  | 
\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse%
 | 
1355  | 
\ my{\isaliteral{5F}{\isacharunderscore}}rule\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline
 | 
|
| 42704 | 1356  | 
\ \ \ \ Attrib{\isaliteral{2E}{\isachardot}}thms\ {\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}fn\ ths\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline
 | 
1357  | 
\ \ \ \ \ \ Thm{\isaliteral{2E}{\isachardot}}rule{\isaliteral{5F}{\isacharunderscore}}attribute\isanewline
 | 
|
1358  | 
\ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}fn\ context{\isaliteral{3A}{\isacharcolon}}\ Context{\isaliteral{2E}{\isachardot}}generic\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ th{\isaliteral{3A}{\isacharcolon}}\ thm\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline
 | 
|
| 40406 | 1359  | 
\ \ \ \ \ \ \ \ \ \ let\ val\ th{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ th\ OF\ ths\isanewline
 | 
| 42936 | 1360  | 
\ \ \ \ \ \ \ \ \ \ in\ th{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}\isanewline
 | 
| 30527 | 1361  | 
\isanewline  | 
| 42704 | 1362  | 
\ \ \isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse%
 | 
| 40406 | 1363  | 
\ my{\isaliteral{5F}{\isacharunderscore}}declaration\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline
 | 
| 42704 | 1364  | 
\ \ \ \ Attrib{\isaliteral{2E}{\isachardot}}thms\ {\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}fn\ ths\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline
 | 
1365  | 
\ \ \ \ \ \ Thm{\isaliteral{2E}{\isachardot}}declaration{\isaliteral{5F}{\isacharunderscore}}attribute\isanewline
 | 
|
1366  | 
\ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}fn\ th{\isaliteral{3A}{\isacharcolon}}\ thm\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ context{\isaliteral{3A}{\isacharcolon}}\ Context{\isaliteral{2E}{\isachardot}}generic\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline
 | 
|
| 40406 | 1367  | 
\ \ \ \ \ \ \ \ \ \ let\ val\ context{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ context\isanewline
 | 
| 42936 | 1368  | 
\ \ \ \ \ \ \ \ \ \ in\ context{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}%
 | 
| 30527 | 1369  | 
\endisatagML  | 
1370  | 
{\isafoldML}%
 | 
|
1371  | 
%  | 
|
1372  | 
\isadelimML  | 
|
1373  | 
%  | 
|
1374  | 
\endisadelimML  | 
|
1375  | 
%  | 
|
| 27042 | 1376  | 
\isamarkupsection{Primitive specification elements%
 | 
1377  | 
}  | 
|
1378  | 
\isamarkuptrue%  | 
|
1379  | 
%  | 
|
1380  | 
\isamarkupsubsection{Type classes and sorts \label{sec:classes}%
 | 
|
1381  | 
}  | 
|
1382  | 
\isamarkuptrue%  | 
|
1383  | 
%  | 
|
1384  | 
\begin{isamarkuptext}%
 | 
|
1385  | 
\begin{matharray}{rcll}
 | 
|
| 40406 | 1386  | 
    \indexdef{}{command}{classes}\hypertarget{command.classes}{\hyperlink{command.classes}{\mbox{\isa{\isacommand{classes}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
1387  | 
    \indexdef{}{command}{classrel}\hypertarget{command.classrel}{\hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\
 | 
|
1388  | 
    \indexdef{}{command}{default\_sort}\hypertarget{command.default-sort}{\hyperlink{command.default-sort}{\mbox{\isa{\isacommand{default{\isaliteral{5F}{\isacharunderscore}}sort}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}}
 | 
|
| 27042 | 1389  | 
  \end{matharray}
 | 
1390  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1391  | 
  \begin{railoutput}
 | 
| 42662 | 1392  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1393  | 
\rail@term{\hyperlink{command.classes}{\mbox{\isa{\isacommand{classes}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1394  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1395  | 
\rail@nont{\hyperlink{syntax.classdecl}{\mbox{\isa{classdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1396  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1397  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1398  | 
\rail@end  | 
| 42662 | 1399  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1400  | 
\rail@term{\hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1401  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1402  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1403  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1404  | 
\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1405  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1406  | 
\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1407  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1408  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1409  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1410  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1411  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1412  | 
\rail@end  | 
| 42662 | 1413  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1414  | 
\rail@term{\hyperlink{command.default-sort}{\mbox{\isa{\isacommand{default{\isaliteral{5F}{\isacharunderscore}}sort}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1415  | 
\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1416  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1417  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1418  | 
|
| 27042 | 1419  | 
|
| 28788 | 1420  | 
  \begin{description}
 | 
| 27042 | 1421  | 
|
| 40406 | 1422  | 
  \item \hyperlink{command.classes}{\mbox{\isa{\isacommand{classes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} declares class
 | 
1423  | 
  \isa{c} to be a subclass of existing classes \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 28788 | 1424  | 
Isabelle implicitly maintains the transitive closure of the class  | 
1425  | 
hierarchy. Cyclic class structures are not permitted.  | 
|
1426  | 
||
| 40406 | 1427  | 
  \item \hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} states subclass
 | 
1428  | 
  relations between existing classes \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}.
 | 
|
| 
37768
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1429  | 
  This is done axiomatically!  The \indexref{}{command}{subclass}\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}} and
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1430  | 
  \indexref{}{command}{instance}\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} commands (see \secref{sec:class}) provide
 | 
| 
 
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
 
haftmann 
parents: 
37115 
diff
changeset
 | 
1431  | 
a way to introduce proven class relations.  | 
| 27042 | 1432  | 
|
| 40406 | 1433  | 
  \item \hyperlink{command.default-sort}{\mbox{\isa{\isacommand{default{\isaliteral{5F}{\isacharunderscore}}sort}}}}~\isa{s} makes sort \isa{s} the
 | 
| 28788 | 1434  | 
new default sort for any type variable that is given explicitly in  | 
1435  | 
the text, but lacks a sort constraint (wrt.\ the current context).  | 
|
1436  | 
Type variables generated by type inference are not affected.  | 
|
| 27042 | 1437  | 
|
| 28788 | 1438  | 
Usually the default sort is only changed when defining a new  | 
1439  | 
object-logic. For example, the default sort in Isabelle/HOL is  | 
|
| 39831 | 1440  | 
  \isa{type}, the class of all HOL types.
 | 
| 27042 | 1441  | 
|
| 28788 | 1442  | 
When merging theories, the default sorts of the parents are  | 
1443  | 
logically intersected, i.e.\ the representations as lists of classes  | 
|
1444  | 
are joined.  | 
|
1445  | 
||
1446  | 
  \end{description}%
 | 
|
| 27042 | 1447  | 
\end{isamarkuptext}%
 | 
1448  | 
\isamarkuptrue%  | 
|
1449  | 
%  | 
|
1450  | 
\isamarkupsubsection{Types and type abbreviations \label{sec:types-pure}%
 | 
|
1451  | 
}  | 
|
1452  | 
\isamarkuptrue%  | 
|
1453  | 
%  | 
|
1454  | 
\begin{isamarkuptext}%
 | 
|
1455  | 
\begin{matharray}{rcll}
 | 
|
| 
41249
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1456  | 
    \indexdef{}{command}{type\_synonym}\hypertarget{command.type-synonym}{\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 40406 | 1457  | 
    \indexdef{}{command}{typedecl}\hypertarget{command.typedecl}{\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
1458  | 
    \indexdef{}{command}{arities}\hypertarget{command.arities}{\hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\
 | 
|
| 27042 | 1459  | 
  \end{matharray}
 | 
1460  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1461  | 
  \begin{railoutput}
 | 
| 42662 | 1462  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1463  | 
\rail@term{\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1464  | 
\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1465  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1466  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1467  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1468  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1469  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1470  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1471  | 
\rail@end  | 
| 42662 | 1472  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1473  | 
\rail@term{\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1474  | 
\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1475  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1476  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1477  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1478  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1479  | 
\rail@end  | 
| 42662 | 1480  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1481  | 
\rail@term{\hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1482  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1483  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1484  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1485  | 
\rail@nont{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1486  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1487  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1488  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1489  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1490  | 
|
| 27042 | 1491  | 
|
| 28788 | 1492  | 
  \begin{description}
 | 
| 27042 | 1493  | 
|
| 
41249
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1494  | 
  \item \hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}
 | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1495  | 
  introduces a \emph{type synonym} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} for the
 | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1496  | 
  existing type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}.  Unlike actual type definitions, as are
 | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1497  | 
available in Isabelle/HOL for example, type synonyms are merely  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1498  | 
syntactic abbreviations without any logical significance.  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
40802 
diff
changeset
 | 
1499  | 
Internally, type synonyms are fully expanded.  | 
| 27042 | 1500  | 
|
| 40406 | 1501  | 
  \item \hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} declares a new
 | 
| 28788 | 1502  | 
  type constructor \isa{t}.  If the object-logic defines a base sort
 | 
1503  | 
  \isa{s}, then the constructor is declared to operate on that, via
 | 
|
| 40406 | 1504  | 
  the axiomatic specification \hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s{\isaliteral{29}{\isacharparenright}}s{\isaliteral{22}{\isachardoublequote}}}.
 | 
| 28788 | 1505  | 
|
| 40406 | 1506  | 
  \item \hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}s{\isaliteral{22}{\isachardoublequote}}} augments
 | 
| 28788 | 1507  | 
Isabelle's order-sorted signature of types by new type constructor  | 
| 
35282
 
8fd9d555d04d
dropped references to old axclass from documentation
 
haftmann 
parents: 
33868 
diff
changeset
 | 
1508  | 
  arities.  This is done axiomatically!  The \indexref{}{command}{instantiation}\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}
 | 
| 
 
8fd9d555d04d
dropped references to old axclass from documentation
 
haftmann 
parents: 
33868 
diff
changeset
 | 
1509  | 
  target (see \secref{sec:class}) provides a way to introduce
 | 
| 28788 | 1510  | 
proven type arities.  | 
1511  | 
||
1512  | 
  \end{description}%
 | 
|
1513  | 
\end{isamarkuptext}%
 | 
|
1514  | 
\isamarkuptrue%  | 
|
1515  | 
%  | 
|
| 27042 | 1516  | 
\isamarkupsubsection{Constants and definitions \label{sec:consts}%
 | 
1517  | 
}  | 
|
1518  | 
\isamarkuptrue%  | 
|
1519  | 
%  | 
|
1520  | 
\begin{isamarkuptext}%
 | 
|
1521  | 
Definitions essentially express abbreviations within the logic. The  | 
|
| 40406 | 1522  | 
  simplest form of a definition is \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}, where \isa{c} is a newly declared constant.  Isabelle also allows derived forms
 | 
| 27042 | 1523  | 
  where the arguments of \isa{c} appear on the left, abbreviating a
 | 
| 40406 | 1524  | 
  prefix of \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}}-abstractions, e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ t{\isaliteral{22}{\isachardoublequote}}} may be
 | 
1525  | 
  written more conveniently as \isa{{\isaliteral{22}{\isachardoublequote}}c\ x\ y\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}.  Moreover,
 | 
|
| 27042 | 1526  | 
definitions may be weakened by adding arbitrary pre-conditions:  | 
| 40406 | 1527  | 
  \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ c\ x\ y\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}.
 | 
| 27042 | 1528  | 
|
1529  | 
\medskip The built-in well-formedness conditions for definitional  | 
|
1530  | 
specifications are:  | 
|
1531  | 
||
1532  | 
  \begin{itemize}
 | 
|
1533  | 
||
1534  | 
\item Arguments (on the left-hand side) must be distinct variables.  | 
|
1535  | 
||
1536  | 
\item All variables on the right-hand side must also appear on the  | 
|
1537  | 
left-hand side.  | 
|
1538  | 
||
1539  | 
\item All type variables on the right-hand side must also appear on  | 
|
| 40406 | 1540  | 
  the left-hand side; this prohibits \isa{{\isaliteral{22}{\isachardoublequote}}{\isadigit{0}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ length\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ list{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} for example.
 | 
| 27042 | 1541  | 
|
1542  | 
\item The definition must not be recursive. Most object-logics  | 
|
1543  | 
provide definitional principles that can be used to express  | 
|
1544  | 
recursion safely.  | 
|
1545  | 
||
1546  | 
  \end{itemize}
 | 
|
1547  | 
||
| 31047 | 1548  | 
The right-hand side of overloaded definitions may mention overloaded constants  | 
| 27042 | 1549  | 
recursively at type instances corresponding to the immediate  | 
| 40406 | 1550  | 
  argument types \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}.  Incomplete
 | 
| 27042 | 1551  | 
specification patterns impose global constraints on all occurrences,  | 
| 40406 | 1552  | 
  e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}d\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}} on the left-hand side means that all
 | 
| 27042 | 1553  | 
corresponding occurrences on some right-hand side need to be an  | 
| 40406 | 1554  | 
  instance of this, general \isa{{\isaliteral{22}{\isachardoublequote}}d\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{22}{\isachardoublequote}}} will be disallowed.
 | 
| 27042 | 1555  | 
|
1556  | 
  \begin{matharray}{rcl}
 | 
|
| 40406 | 1557  | 
    \indexdef{}{command}{consts}\hypertarget{command.consts}{\hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
1558  | 
    \indexdef{}{command}{defs}\hypertarget{command.defs}{\hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27042 | 1559  | 
  \end{matharray}
 | 
1560  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1561  | 
  \begin{railoutput}
 | 
| 42662 | 1562  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1563  | 
\rail@term{\hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1564  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1565  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1566  | 
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1567  | 
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1568  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1569  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1570  | 
\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1571  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1572  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1573  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1574  | 
\rail@end  | 
| 42704 | 1575  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1576  | 
\rail@term{\hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1577  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1578  | 
\rail@nextbar{1}
 | 
| 42704 | 1579  | 
\rail@nont{\isa{opt}}[]
 | 
1580  | 
\rail@endbar  | 
|
1581  | 
\rail@plus  | 
|
1582  | 
\rail@nont{\hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}}[]
 | 
|
1583  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
|
1584  | 
\rail@nextplus{1}
 | 
|
1585  | 
\rail@endplus  | 
|
1586  | 
\rail@end  | 
|
1587  | 
\rail@begin{2}{\isa{opt}}
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1588  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1589  | 
\rail@bar  | 
| 42704 | 1590  | 
\rail@nextbar{1}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1591  | 
\rail@term{\isa{\isakeyword{unchecked}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1592  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1593  | 
\rail@bar  | 
| 42704 | 1594  | 
\rail@nextbar{1}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1595  | 
\rail@term{\isa{\isakeyword{overloaded}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1596  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1597  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1598  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1599  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1600  | 
|
| 27042 | 1601  | 
|
| 28788 | 1602  | 
  \begin{description}
 | 
| 27042 | 1603  | 
|
| 40406 | 1604  | 
  \item \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{22}{\isachardoublequote}}} declares constant \isa{c} to have any instance of type scheme \isa{{\isaliteral{5C3C7369676D613E}{\isasymsigma}}}.  The optional
 | 
| 28788 | 1605  | 
mixfix annotations may attach concrete syntax to the constants  | 
1606  | 
declared.  | 
|
| 27042 | 1607  | 
|
| 40406 | 1608  | 
  \item \hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ eqn{\isaliteral{22}{\isachardoublequote}}} introduces \isa{eqn}
 | 
| 27042 | 1609  | 
as a definitional axiom for some existing constant.  | 
1610  | 
||
| 40406 | 1611  | 
  The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}unchecked{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option disables global dependency checks
 | 
| 27042 | 1612  | 
for this definition, which is occasionally useful for exotic  | 
1613  | 
overloading. It is at the discretion of the user to avoid malformed  | 
|
1614  | 
theory specifications!  | 
|
1615  | 
||
| 40406 | 1616  | 
  The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}overloaded{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option declares definitions to be
 | 
| 27042 | 1617  | 
potentially overloaded. Unless this option is given, a warning  | 
1618  | 
message would be issued for any definitional equation with a more  | 
|
1619  | 
special type than that of the corresponding constant declaration.  | 
|
1620  | 
||
| 28788 | 1621  | 
  \end{description}%
 | 
| 27042 | 1622  | 
\end{isamarkuptext}%
 | 
1623  | 
\isamarkuptrue%  | 
|
1624  | 
%  | 
|
1625  | 
\isamarkupsection{Axioms and theorems \label{sec:axms-thms}%
 | 
|
1626  | 
}  | 
|
1627  | 
\isamarkuptrue%  | 
|
1628  | 
%  | 
|
1629  | 
\begin{isamarkuptext}%
 | 
|
1630  | 
\begin{matharray}{rcll}
 | 
|
| 40406 | 1631  | 
    \indexdef{}{command}{axioms}\hypertarget{command.axioms}{\hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\
 | 
1632  | 
    \indexdef{}{command}{lemmas}\hypertarget{command.lemmas}{\hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1633  | 
    \indexdef{}{command}{theorems}\hypertarget{command.theorems}{\hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27042 | 1634  | 
  \end{matharray}
 | 
1635  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1636  | 
  \begin{railoutput}
 | 
| 42662 | 1637  | 
\rail@begin{2}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1638  | 
\rail@term{\hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1639  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1640  | 
\rail@nont{\hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1641  | 
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1642  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1643  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1644  | 
\rail@end  | 
| 42662 | 1645  | 
\rail@begin{3}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1646  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1647  | 
\rail@term{\hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1648  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1649  | 
\rail@term{\hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1650  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1651  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1652  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1653  | 
\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1654  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1655  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1656  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1657  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1658  | 
\rail@nont{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1659  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1660  | 
\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1661  | 
\rail@nextplus{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1662  | 
\rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1663  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1664  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1665  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1666  | 
|
| 27042 | 1667  | 
|
| 28788 | 1668  | 
  \begin{description}
 | 
| 27042 | 1669  | 
|
| 40406 | 1670  | 
  \item \hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} introduces arbitrary
 | 
| 27042 | 1671  | 
statements as axioms of the meta-logic. In fact, axioms are  | 
1672  | 
``axiomatic theorems'', and may be referred later just as any other  | 
|
1673  | 
theorem.  | 
|
1674  | 
||
1675  | 
Axioms are usually only introduced when declaring new logical  | 
|
1676  | 
systems. Everyday work is typically done the hard way, with proper  | 
|
1677  | 
definitions and proven theorems.  | 
|
1678  | 
||
| 40406 | 1679  | 
  \item \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3D}{\isacharequal}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} retrieves and stores
 | 
| 28788 | 1680  | 
existing facts in the theory context, or the specified target  | 
1681  | 
  context (see also \secref{sec:target}).  Typical applications would
 | 
|
1682  | 
also involve attributes, to declare Simplifier rules, for example.  | 
|
| 27042 | 1683  | 
|
| 28788 | 1684  | 
  \item \hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}} is essentially the same as \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}, but marks the result as a different kind of facts.
 | 
| 27042 | 1685  | 
|
| 28788 | 1686  | 
  \end{description}%
 | 
| 27042 | 1687  | 
\end{isamarkuptext}%
 | 
1688  | 
\isamarkuptrue%  | 
|
1689  | 
%  | 
|
1690  | 
\isamarkupsection{Oracles%
 | 
|
1691  | 
}  | 
|
1692  | 
\isamarkuptrue%  | 
|
1693  | 
%  | 
|
1694  | 
\begin{isamarkuptext}%
 | 
|
| 28788 | 1695  | 
Oracles allow Isabelle to take advantage of external reasoners  | 
1696  | 
such as arithmetic decision procedures, model checkers, fast  | 
|
1697  | 
tautology checkers or computer algebra systems. Invoked as an  | 
|
1698  | 
oracle, an external reasoner can create arbitrary Isabelle theorems.  | 
|
| 27042 | 1699  | 
|
| 28788 | 1700  | 
It is the responsibility of the user to ensure that the external  | 
1701  | 
reasoner is as trustworthy as the application requires. Another  | 
|
1702  | 
typical source of errors is the linkup between Isabelle and the  | 
|
1703  | 
external tool, not just its concrete implementation, but also the  | 
|
1704  | 
required translation between two different logical environments.  | 
|
1705  | 
||
1706  | 
Isabelle merely guarantees well-formedness of the propositions being  | 
|
1707  | 
asserted, and records within the internal derivation object how  | 
|
1708  | 
presumed theorems depend on unproven suppositions.  | 
|
1709  | 
||
| 40240 | 1710  | 
  \begin{matharray}{rcll}
 | 
| 40406 | 1711  | 
    \indexdef{}{command}{oracle}\hypertarget{command.oracle}{\hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\
 | 
| 28788 | 1712  | 
  \end{matharray}
 | 
| 27042 | 1713  | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1714  | 
  \begin{railoutput}
 | 
| 42662 | 1715  | 
\rail@begin{1}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1716  | 
\rail@term{\hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1717  | 
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1718  | 
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1719  | 
\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1720  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1721  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1722  | 
|
| 27042 | 1723  | 
|
| 28788 | 1724  | 
  \begin{description}
 | 
| 27042 | 1725  | 
|
| 40406 | 1726  | 
  \item \hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ text{\isaliteral{22}{\isachardoublequote}}} turns the given ML
 | 
1727  | 
  expression \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} of type \verb|'a -> cterm| into an
 | 
|
| 28291 | 1728  | 
ML function of type \verb|'a -> thm|, which is bound to the  | 
| 28788 | 1729  | 
global identifier \verb|name|. This acts like an infinitary  | 
1730  | 
specification of axioms! Invoking the oracle only works within the  | 
|
1731  | 
scope of the resulting theory.  | 
|
| 27042 | 1732  | 
|
| 28788 | 1733  | 
  \end{description}
 | 
1734  | 
||
| 40802 | 1735  | 
See \verb|~~/src/HOL/ex/Iff_Oracle.thy| for a worked example of  | 
| 28788 | 1736  | 
defining a new primitive rule as oracle, and turning it into a proof  | 
1737  | 
method.%  | 
|
| 27042 | 1738  | 
\end{isamarkuptext}%
 | 
1739  | 
\isamarkuptrue%  | 
|
1740  | 
%  | 
|
1741  | 
\isamarkupsection{Name spaces%
 | 
|
1742  | 
}  | 
|
1743  | 
\isamarkuptrue%  | 
|
1744  | 
%  | 
|
1745  | 
\begin{isamarkuptext}%
 | 
|
1746  | 
\begin{matharray}{rcl}
 | 
|
| 40406 | 1747  | 
    \indexdef{}{command}{hide\_class}\hypertarget{command.hide-class}{\hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
1748  | 
    \indexdef{}{command}{hide\_type}\hypertarget{command.hide-type}{\hyperlink{command.hide-type}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1749  | 
    \indexdef{}{command}{hide\_const}\hypertarget{command.hide-const}{\hyperlink{command.hide-const}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}const}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
1750  | 
    \indexdef{}{command}{hide\_fact}\hypertarget{command.hide-fact}{\hyperlink{command.hide-fact}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}fact}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
|
| 27042 | 1751  | 
  \end{matharray}
 | 
1752  | 
||
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1753  | 
  \begin{railoutput}
 | 
| 42662 | 1754  | 
\rail@begin{4}{}
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1755  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1756  | 
\rail@nont{\hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1757  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1758  | 
\rail@nont{\hyperlink{command.hide-type}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}type}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1759  | 
\rail@nextbar{2}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1760  | 
\rail@nont{\hyperlink{command.hide-const}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}const}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1761  | 
\rail@nextbar{3}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1762  | 
\rail@nont{\hyperlink{command.hide-fact}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}fact}}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1763  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1764  | 
\rail@bar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1765  | 
\rail@nextbar{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1766  | 
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1767  | 
\rail@term{\isa{\isakeyword{open}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1768  | 
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1769  | 
\rail@endbar  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1770  | 
\rail@plus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1771  | 
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1772  | 
\rail@nextplus{1}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1773  | 
\rail@endplus  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1774  | 
\rail@end  | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1775  | 
\end{railoutput}
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
1776  | 
|
| 27042 | 1777  | 
|
1778  | 
Isabelle organizes any kind of name declarations (of types,  | 
|
1779  | 
constants, theorems etc.) by separate hierarchically structured name  | 
|
1780  | 
spaces. Normally the user does not have to control the behavior of  | 
|
1781  | 
name spaces by hand, yet the following commands provide some way to  | 
|
1782  | 
do so.  | 
|
1783  | 
||
| 28788 | 1784  | 
  \begin{description}
 | 
| 27042 | 1785  | 
|
| 40406 | 1786  | 
  \item \hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}~\isa{names} fully removes class
 | 
1787  | 
  declarations from a given name space; with the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}open{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}
 | 
|
| 
39977
 
c9cbc16e93ce
do not mention unqualified names, now that 'global' and 'local' are gone
 
krauss 
parents: 
39280 
diff
changeset
 | 
1788  | 
option, only the base name is hidden.  | 
| 
36177
 
8e0770d2e499
separate commands 'hide_class', 'hide_type', 'hide_const', 'hide_fact';
 
wenzelm 
parents: 
35681 
diff
changeset
 | 
1789  | 
|
| 27042 | 1790  | 
Note that hiding name space accesses has no impact on logical  | 
| 28788 | 1791  | 
declarations --- they remain valid internally. Entities that are no  | 
| 27042 | 1792  | 
longer accessible to the user are printed with the special qualifier  | 
| 40406 | 1793  | 
  ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' prefixed to the full internal name.
 | 
| 27042 | 1794  | 
|
| 40406 | 1795  | 
  \item \hyperlink{command.hide-type}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}type}}}}, \hyperlink{command.hide-const}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}const}}}}, and \hyperlink{command.hide-fact}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}fact}}}} are similar to \hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}, but hide types,
 | 
| 
36177
 
8e0770d2e499
separate commands 'hide_class', 'hide_type', 'hide_const', 'hide_fact';
 
wenzelm 
parents: 
35681 
diff
changeset
 | 
1796  | 
constants, and facts, respectively.  | 
| 
 
8e0770d2e499
separate commands 'hide_class', 'hide_type', 'hide_const', 'hide_fact';
 
wenzelm 
parents: 
35681 
diff
changeset
 | 
1797  | 
|
| 28788 | 1798  | 
  \end{description}%
 | 
| 27042 | 1799  | 
\end{isamarkuptext}%
 | 
1800  | 
\isamarkuptrue%  | 
|
1801  | 
%  | 
|
| 26869 | 1802  | 
\isadelimtheory  | 
1803  | 
%  | 
|
1804  | 
\endisadelimtheory  | 
|
1805  | 
%  | 
|
1806  | 
\isatagtheory  | 
|
1807  | 
\isacommand{end}\isamarkupfalse%
 | 
|
1808  | 
%  | 
|
1809  | 
\endisatagtheory  | 
|
1810  | 
{\isafoldtheory}%
 | 
|
1811  | 
%  | 
|
1812  | 
\isadelimtheory  | 
|
1813  | 
%  | 
|
1814  | 
\endisadelimtheory  | 
|
1815  | 
\isanewline  | 
|
1816  | 
\end{isabellebody}%
 | 
|
1817  | 
%%% Local Variables:  | 
|
1818  | 
%%% mode: latex  | 
|
1819  | 
%%% TeX-master: "root"  | 
|
1820  | 
%%% End:  |