author | wenzelm |
Sat, 05 Oct 2024 14:58:36 +0200 | |
changeset 81116 | 0fb1e2dd4122 |
parent 81113 | 6fefd6c602fa |
child 81514 | 98cb63b447c6 |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
2 |
||
26869 | 3 |
theory Spec |
63531 | 4 |
imports Main Base |
26869 | 5 |
begin |
6 |
||
58618 | 7 |
chapter \<open>Specifications\<close> |
26869 | 8 |
|
61664 | 9 |
text \<open> |
10 |
The Isabelle/Isar theory format integrates specifications and proofs, with |
|
11 |
support for interactive development by continuous document editing. There is |
|
12 |
a separate document preparation system (see \chref{ch:document-prep}), for |
|
13 |
typesetting formal developments together with informal text. The resulting |
|
14 |
hyper-linked PDF documents can be used both for WWW presentation and printed |
|
15 |
copies. |
|
29745 | 16 |
|
61664 | 17 |
The Isar proof language (see \chref{ch:proofs}) is embedded into the theory |
18 |
language as a proper sub-language. Proof mode is entered by stating some |
|
61665 | 19 |
\<^theory_text>\<open>theorem\<close> or \<^theory_text>\<open>lemma\<close> at the theory level, and left again with the final |
20 |
conclusion (e.g.\ via \<^theory_text>\<open>qed\<close>). |
|
58618 | 21 |
\<close> |
29745 | 22 |
|
23 |
||
58618 | 24 |
section \<open>Defining theories \label{sec:begin-thy}\<close> |
26870 | 25 |
|
58618 | 26 |
text \<open> |
26870 | 27 |
\begin{matharray}{rcl} |
61493 | 28 |
@{command_def "theory"} & : & \<open>toplevel \<rightarrow> theory\<close> \\ |
29 |
@{command_def (global) "end"} & : & \<open>theory \<rightarrow> toplevel\<close> \\ |
|
30 |
@{command_def "thy_deps"}\<open>\<^sup>*\<close> & : & \<open>theory \<rightarrow>\<close> \\ |
|
26870 | 31 |
\end{matharray} |
32 |
||
59781 | 33 |
Isabelle/Isar theories are defined via theory files, which consist of an |
34 |
outermost sequence of definition--statement--proof elements. Some |
|
61665 | 35 |
definitions are self-sufficient (e.g.\ \<^theory_text>\<open>fun\<close> in Isabelle/HOL), with |
61664 | 36 |
foundational proofs performed internally. Other definitions require an |
61665 | 37 |
explicit proof as justification (e.g.\ \<^theory_text>\<open>function\<close> and \<^theory_text>\<open>termination\<close> in |
38 |
Isabelle/HOL). Plain statements like \<^theory_text>\<open>theorem\<close> or \<^theory_text>\<open>lemma\<close> are merely a |
|
39 |
special case of that, defining a theorem from a given proposition and its |
|
40 |
proof. |
|
26870 | 41 |
|
61664 | 42 |
The theory body may be sub-structured by means of \<^emph>\<open>local theory targets\<close>, |
61706 | 43 |
such as \<^theory_text>\<open>locale\<close> and \<^theory_text>\<open>class\<close>. It is also possible to use \<^theory_text>\<open>context begin \<dots> |
44 |
end\<close> blocks to delimited a local theory context: a \<^emph>\<open>named context\<close> to |
|
45 |
augment a locale or class specification, or an \<^emph>\<open>unnamed context\<close> to refer |
|
46 |
to local parameters and assumptions that are discharged later. See |
|
47 |
\secref{sec:target} for more details. |
|
28745 | 48 |
|
61421 | 49 |
\<^medskip> |
61665 | 50 |
A theory is commenced by the \<^theory_text>\<open>theory\<close> command, which indicates imports of |
51 |
previous theories, according to an acyclic foundational order. Before the |
|
52 |
initial \<^theory_text>\<open>theory\<close> command, there may be optional document header material |
|
53 |
(like \<^theory_text>\<open>section\<close> or \<^theory_text>\<open>text\<close>, see \secref{sec:markup}). The document header |
|
54 |
is outside of the formal theory context, though. |
|
59781 | 55 |
|
61664 | 56 |
A theory is concluded by a final @{command (global) "end"} command, one that |
57 |
does not belong to a local theory target. No further commands may follow |
|
58 |
such a global @{command (global) "end"}. |
|
26870 | 59 |
|
69597 | 60 |
\<^rail>\<open> |
66916 | 61 |
@@{command theory} @{syntax system_name} |
62 |
@'imports' (@{syntax system_name} +) \<newline> |
|
63580 | 63 |
keywords? abbrevs? @'begin' |
47114
7c9e31ffcd9e
updated theory header syntax and related details;
wenzelm
parents:
46999
diff
changeset
|
64 |
; |
50128
599c935aac82
alternative completion for outer syntax keywords;
wenzelm
parents:
49243
diff
changeset
|
65 |
keywords: @'keywords' (keyword_decls + @'and') |
599c935aac82
alternative completion for outer syntax keywords;
wenzelm
parents:
49243
diff
changeset
|
66 |
; |
63579 | 67 |
keyword_decls: (@{syntax string} +) ('::' @{syntax name} @{syntax tags})? |
68 |
; |
|
67722
012f1e8a1209
allow multiple entries of and_list (on both sides);
wenzelm
parents:
67013
diff
changeset
|
69 |
abbrevs: @'abbrevs' (((text+) '=' (text+)) + @'and') |
60093 | 70 |
; |
71 |
@@{command thy_deps} (thy_bounds thy_bounds?)? |
|
72 |
; |
|
73 |
thy_bounds: @{syntax name} | '(' (@{syntax name} + @'|') ')' |
|
69597 | 74 |
\<close> |
26870 | 75 |
|
61665 | 76 |
\<^descr> \<^theory_text>\<open>theory A imports B\<^sub>1 \<dots> B\<^sub>n begin\<close> starts a new theory \<open>A\<close> based on the |
77 |
merge of existing theories \<open>B\<^sub>1 \<dots> B\<^sub>n\<close>. Due to the possibility to import |
|
78 |
more than one ancestor, the resulting theory structure of an Isabelle |
|
79 |
session forms a directed acyclic graph (DAG). Isabelle takes care that |
|
80 |
sources contributing to the development graph are always up-to-date: changed |
|
81 |
files are automatically rechecked whenever a theory header specification is |
|
82 |
processed. |
|
47114
7c9e31ffcd9e
updated theory header syntax and related details;
wenzelm
parents:
46999
diff
changeset
|
83 |
|
59781 | 84 |
Empty imports are only allowed in the bootstrap process of the special |
69597 | 85 |
theory \<^theory>\<open>Pure\<close>, which is the start of any other formal development |
61664 | 86 |
based on Isabelle. Regular user theories usually refer to some more complex |
69597 | 87 |
entry point, such as theory \<^theory>\<open>Main\<close> in Isabelle/HOL. |
59781 | 88 |
|
63579 | 89 |
The @{keyword_def "keywords"} specification declares outer syntax |
61664 | 90 |
(\chref{ch:outer-syntax}) that is introduced in this theory later on (rare |
91 |
in end-user applications). Both minor keywords and major keywords of the |
|
92 |
Isar command language need to be specified, in order to make parsing of |
|
93 |
proof documents work properly. Command keywords need to be classified |
|
94 |
according to their structural role in the formal text. Examples may be seen |
|
95 |
in Isabelle/HOL sources itself, such as @{keyword "keywords"}~\<^verbatim>\<open>"typedef"\<close> |
|
69913 | 96 |
\<open>:: thy_goal_defn\<close> or @{keyword "keywords"}~\<^verbatim>\<open>"datatype"\<close> \<open>:: thy_defn\<close> for |
97 |
theory-level definitions with and without proof, respectively. Additional |
|
61664 | 98 |
@{syntax tags} provide defaults for document preparation |
69962
82e945d472d5
documentation of document markers and re-interpreted command tags;
wenzelm
parents:
69913
diff
changeset
|
99 |
(\secref{sec:document-markers}). |
50128
599c935aac82
alternative completion for outer syntax keywords;
wenzelm
parents:
49243
diff
changeset
|
100 |
|
63579 | 101 |
The @{keyword_def "abbrevs"} specification declares additional abbreviations |
102 |
for syntactic completion. The default for a new keyword is just its name, |
|
103 |
but completion may be avoided by defining @{keyword "abbrevs"} with empty |
|
104 |
text. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
105 |
|
61664 | 106 |
\<^descr> @{command (global) "end"} concludes the current theory definition. Note |
61665 | 107 |
that some other commands, e.g.\ local theory targets \<^theory_text>\<open>locale\<close> or \<^theory_text>\<open>class\<close> |
108 |
may involve a \<^theory_text>\<open>begin\<close> that needs to be matched by @{command (local) "end"}, |
|
109 |
according to the usual rules for nested blocks. |
|
27040 | 110 |
|
61665 | 111 |
\<^descr> \<^theory_text>\<open>thy_deps\<close> visualizes the theory hierarchy as a directed acyclic graph. |
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
112 |
By default, all imported theories are shown. This may be restricted by |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
113 |
specifying bounds wrt. the theory inclusion relation. |
58618 | 114 |
\<close> |
27040 | 115 |
|
116 |
||
58618 | 117 |
section \<open>Local theory targets \label{sec:target}\<close> |
27040 | 118 |
|
58618 | 119 |
text \<open> |
47483 | 120 |
\begin{matharray}{rcll} |
61493 | 121 |
@{command_def "context"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
122 |
@{command_def (local) "end"} & : & \<open>local_theory \<rightarrow> theory\<close> \\ |
|
59926 | 123 |
@{keyword_def "private"} \\ |
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59939
diff
changeset
|
124 |
@{keyword_def "qualified"} \\ |
47483 | 125 |
\end{matharray} |
126 |
||
61664 | 127 |
A local theory target is a specification context that is managed separately |
128 |
within the enclosing theory. Contexts may introduce parameters (fixed |
|
129 |
variables) and assumptions (hypotheses). Definitions and theorems depending |
|
130 |
on the context may be added incrementally later on. |
|
47482
a83b25e5bad3
some coverage of unnamed contexts, which can be nested within other targets;
wenzelm
parents:
47114
diff
changeset
|
131 |
|
61664 | 132 |
\<^emph>\<open>Named contexts\<close> refer to locales (cf.\ \secref{sec:locale}) or type |
133 |
classes (cf.\ \secref{sec:class}); the name ``\<open>-\<close>'' signifies the global |
|
134 |
theory context. |
|
47482
a83b25e5bad3
some coverage of unnamed contexts, which can be nested within other targets;
wenzelm
parents:
47114
diff
changeset
|
135 |
|
61664 | 136 |
\<^emph>\<open>Unnamed contexts\<close> may introduce additional parameters and assumptions, and |
137 |
results produced in the context are generalized accordingly. Such auxiliary |
|
61665 | 138 |
contexts may be nested within other targets, like \<^theory_text>\<open>locale\<close>, \<^theory_text>\<open>class\<close>, |
139 |
\<^theory_text>\<open>instantiation\<close>, \<^theory_text>\<open>overloading\<close>. |
|
27040 | 140 |
|
69597 | 141 |
\<^rail>\<open> |
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
142 |
@@{command context} @{syntax name} @{syntax_ref "opening"}? @'begin' |
27040 | 143 |
; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
144 |
@@{command context} @{syntax_ref "includes"}? (@{syntax context_elem} * ) @'begin' |
47484 | 145 |
; |
62969 | 146 |
@{syntax_def target}: '(' @'in' @{syntax name} ')' |
69597 | 147 |
\<close> |
27040 | 148 |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
149 |
\<^descr> \<^theory_text>\<open>context c bundles begin\<close> opens a named context, by recommencing an existing |
61665 | 150 |
locale or class \<open>c\<close>. Note that locale and class definitions allow to include |
151 |
the \<^theory_text>\<open>begin\<close> keyword as well, in order to continue the local theory |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
152 |
immediately after the initial specification. Optionally given |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
153 |
\<open>bundles\<close> only take effect in the surface context within the \<^theory_text>\<open>begin\<close> / |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
154 |
\<^theory_text>\<open>end\<close> block. |
47482
a83b25e5bad3
some coverage of unnamed contexts, which can be nested within other targets;
wenzelm
parents:
47114
diff
changeset
|
155 |
|
61665 | 156 |
\<^descr> \<^theory_text>\<open>context bundles elements begin\<close> opens an unnamed context, by extending |
157 |
the enclosing global or local theory target by the given declaration bundles |
|
158 |
(\secref{sec:bundle}) and context elements (\<^theory_text>\<open>fixes\<close>, \<^theory_text>\<open>assumes\<close> etc.). This |
|
159 |
means any results stemming from definitions and proofs in the extended |
|
160 |
context will be exported into the enclosing target by lifting over extra |
|
161 |
parameters and premises. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
162 |
|
61664 | 163 |
\<^descr> @{command (local) "end"} concludes the current local theory, according to |
164 |
the nesting of contexts. Note that a global @{command (global) "end"} has a |
|
165 |
different meaning: it concludes the theory itself (\secref{sec:begin-thy}). |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
166 |
|
61665 | 167 |
\<^descr> \<^theory_text>\<open>private\<close> or \<^theory_text>\<open>qualified\<close> may be given as modifiers before any local |
168 |
theory command. This restricts name space accesses to the local scope, as |
|
169 |
determined by the enclosing \<^theory_text>\<open>context begin \<dots> end\<close> block. Outside its scope, |
|
170 |
a \<^theory_text>\<open>private\<close> name is inaccessible, and a \<^theory_text>\<open>qualified\<close> name is only |
|
61664 | 171 |
accessible with some qualification. |
59939
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents:
59926
diff
changeset
|
172 |
|
61665 | 173 |
Neither a global \<^theory_text>\<open>theory\<close> nor a \<^theory_text>\<open>locale\<close> target provides a local scope by |
174 |
itself: an extra unnamed context is required to use \<^theory_text>\<open>private\<close> or |
|
175 |
\<^theory_text>\<open>qualified\<close> here. |
|
59926 | 176 |
|
61664 | 177 |
\<^descr> \<open>(\<close>@{keyword_def "in"}~\<open>c)\<close> given after any local theory command specifies |
61665 | 178 |
an immediate target, e.g.\ ``\<^theory_text>\<open>definition (in c)\<close>'' or |
179 |
``\<^theory_text>\<open>theorem (in c)\<close>''. This works both in a local or global theory context; |
|
180 |
the current target context will be suspended for this command only. Note |
|
181 |
that ``\<^theory_text>\<open>(in -)\<close>'' will always produce a global result independently of the |
|
182 |
current target context. |
|
27040 | 183 |
|
184 |
||
61664 | 185 |
Any specification element that operates on \<open>local_theory\<close> according to this |
61665 | 186 |
manual implicitly allows the above target syntax \<^theory_text>\<open>(in c)\<close>, but individual |
187 |
syntax diagrams omit that aspect for clarity. |
|
59783
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
wenzelm
parents:
59781
diff
changeset
|
188 |
|
61421 | 189 |
\<^medskip> |
61664 | 190 |
The exact meaning of results produced within a local theory context depends |
191 |
on the underlying target infrastructure (locale, type class etc.). The |
|
192 |
general idea is as follows, considering a context named \<open>c\<close> with parameter |
|
193 |
\<open>x\<close> and assumption \<open>A[x]\<close>. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
194 |
|
61664 | 195 |
Definitions are exported by introducing a global version with additional |
196 |
arguments; a syntactic abbreviation links the long form with the abstract |
|
197 |
version of the target context. For example, \<open>a \<equiv> t[x]\<close> becomes \<open>c.a ?x \<equiv> |
|
198 |
t[?x]\<close> at the theory level (for arbitrary \<open>?x\<close>), together with a local |
|
71567
9a29e883a934
tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents:
71166
diff
changeset
|
199 |
abbreviation \<open>a \<equiv> c.a x\<close> in the target context (for the fixed parameter |
61664 | 200 |
\<open>x\<close>). |
27040 | 201 |
|
61664 | 202 |
Theorems are exported by discharging the assumptions and generalizing the |
203 |
parameters of the context. For example, \<open>a: B[x]\<close> becomes \<open>c.a: A[?x] \<Longrightarrow> |
|
204 |
B[?x]\<close>, again for arbitrary \<open>?x\<close>. |
|
59783
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
wenzelm
parents:
59781
diff
changeset
|
205 |
\<close> |
27040 | 206 |
|
207 |
||
58618 | 208 |
section \<open>Bundled declarations \label{sec:bundle}\<close> |
47484 | 209 |
|
58618 | 210 |
text \<open> |
47484 | 211 |
\begin{matharray}{rcl} |
61493 | 212 |
@{command_def "bundle"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
63273 | 213 |
@{command "bundle"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
81113 | 214 |
@{command_def "unbundle"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
61494 | 215 |
@{command_def "print_bundles"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
61493 | 216 |
@{command_def "include"} & : & \<open>proof(state) \<rightarrow> proof(state)\<close> \\ |
217 |
@{command_def "including"} & : & \<open>proof(prove) \<rightarrow> proof(prove)\<close> \\ |
|
47484 | 218 |
@{keyword_def "includes"} & : & syntax \\ |
81113 | 219 |
@{keyword_def "opening"} & : & syntax \\ |
47484 | 220 |
\end{matharray} |
221 |
||
222 |
The outer syntax of fact expressions (\secref{sec:syn-att}) involves |
|
61664 | 223 |
theorems and attributes, which are evaluated in the context and applied to |
224 |
it. Attributes may declare theorems to the context, as in \<open>this_rule [intro] |
|
225 |
that_rule [elim]\<close> for example. Configuration options (\secref{sec:config}) |
|
226 |
are special declaration attributes that operate on the context without a |
|
227 |
theorem, as in \<open>[[show_types = false]]\<close> for example. |
|
47484 | 228 |
|
61664 | 229 |
Expressions of this form may be defined as \<^emph>\<open>bundled declarations\<close> in the |
230 |
context, and included in other situations later on. Including declaration |
|
231 |
bundles augments a local context casually without logical dependencies, |
|
232 |
which is in contrast to locales and locale interpretation |
|
233 |
(\secref{sec:locale}). |
|
47484 | 234 |
|
69597 | 235 |
\<^rail>\<open> |
81106 | 236 |
(@@{command bundle} | @@{command open_bundle}) @{syntax name} \<newline> |
63273 | 237 |
( '=' @{syntax thms} @{syntax for_fixes} | @'begin') |
63272 | 238 |
; |
81113 | 239 |
@@{command unbundle} @{syntax bundles} |
240 |
; |
|
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
241 |
@@{command print_bundles} ('!'?) |
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
242 |
; |
81113 | 243 |
(@@{command include} | @@{command including}) @{syntax bundles} |
47484 | 244 |
; |
81113 | 245 |
@{syntax_def "includes"}: @'includes' @{syntax bundles} |
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
246 |
; |
81113 | 247 |
@{syntax_def "opening"}: @'opening' @{syntax bundles} |
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
248 |
; |
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
249 |
@{syntax bundles}: ('no')? @{syntax name} + @'and' |
69597 | 250 |
\<close> |
47484 | 251 |
|
61665 | 252 |
\<^descr> \<^theory_text>\<open>bundle b = decls\<close> defines a bundle of declarations in the current |
253 |
context. The RHS is similar to the one of the \<^theory_text>\<open>declare\<close> command. Bundles |
|
254 |
defined in local theory targets are subject to transformations via |
|
255 |
morphisms, when moved into different application contexts; this works |
|
256 |
analogously to any other local theory specification. |
|
47484 | 257 |
|
63273 | 258 |
\<^descr> \<^theory_text>\<open>bundle b begin body end\<close> defines a bundle of declarations from the |
259 |
\<open>body\<close> of local theory specifications. It may consist of commands that are |
|
260 |
technically equivalent to \<^theory_text>\<open>declare\<close> or \<^theory_text>\<open>declaration\<close>, which also includes |
|
261 |
\<^theory_text>\<open>notation\<close>, for example. Named fact declarations like ``\<^theory_text>\<open>lemmas a [simp] = |
|
262 |
b\<close>'' or ``\<^theory_text>\<open>lemma a [simp]: B \<proof>\<close>'' are also admitted, but the name |
|
263 |
bindings are not recorded in the bundle. |
|
63272 | 264 |
|
81106 | 265 |
\<^descr> \<^theory_text>\<open>open_bundle b\<close> is like \<^theory_text>\<open>bundle b\<close> followed by \<^theory_text>\<open>unbundle b\<close>, so its |
81113 | 266 |
declarations are activated immediately, but also named for later re-use. |
267 |
||
268 |
\<^descr> \<^theory_text>\<open>unbundle \<^vec>b\<close> activates the declarations from the given bundles in |
|
269 |
the current local theory context. This is analogous to \<^theory_text>\<open>lemmas\<close> |
|
270 |
(\secref{sec:theorems}) with the expanded bundles. |
|
81106 | 271 |
|
61665 | 272 |
\<^descr> \<^theory_text>\<open>print_bundles\<close> prints the named bundles that are available in the |
273 |
current context; the ``\<open>!\<close>'' option indicates extra verbosity. |
|
47484 | 274 |
|
81113 | 275 |
\<^descr> \<^theory_text>\<open>include \<^vec>b\<close> activates the declarations from the given bundles in a |
276 |
proof body (forward mode). This is analogous to \<^theory_text>\<open>note\<close> |
|
277 |
(\secref{sec:proof-facts}) with the expanded bundles. |
|
278 |
||
279 |
\<^descr> \<^theory_text>\<open>including \<^vec>b\<close> is similar to \<^theory_text>\<open>include\<close>, but works in proof |
|
280 |
refinement (backward mode). This is analogous to \<^theory_text>\<open>using\<close> |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
281 |
(\secref{sec:proof-facts}) with the expanded bundles. |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
282 |
|
81113 | 283 |
\<^descr> \<^theory_text>\<open>includes \<^vec>b\<close> is similar to \<^theory_text>\<open>include\<close>, but applies to a confined |
284 |
specification context: unnamed \<^theory_text>\<open>context\<close>s and long statements of |
|
285 |
\<^theory_text>\<open>theorem\<close>. |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
286 |
|
81113 | 287 |
\<^descr> \<^theory_text>\<open>opening \<^vec>b\<close> is similar to \<^theory_text>\<open>includes\<close>, but applies to a named |
288 |
specification context: \<^theory_text>\<open>locale\<close>s, \<^theory_text>\<open>class\<close>es and named \<^theory_text>\<open>context\<close>s. The |
|
289 |
effect is confined to the surface context within the specification block |
|
290 |
itself and the corresponding \<^theory_text>\<open>begin\<close> / \<^theory_text>\<open>end\<close> block. |
|
47484 | 291 |
|
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
292 |
\<^descr> Bundle names may be prefixed by the reserved word \<^verbatim>\<open>no\<close> to indicate that |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
293 |
the polarity of certain declaration commands should be inverted, notably: |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
294 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
295 |
\<^item> @{command syntax} versus @{command no_syntax} |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
296 |
\<^item> @{command notation} versus @{command no_notation} |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
297 |
\<^item> @{command type_notation} versus @{command no_type_notation} |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
298 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
299 |
This also works recursively for the @{command unbundle} command as |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
300 |
declaration inside a @{command bundle} definition. |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
301 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81113
diff
changeset
|
302 |
|
61664 | 303 |
Here is an artificial example of bundling various configuration options: |
61458 | 304 |
\<close> |
47484 | 305 |
|
59905 | 306 |
(*<*)experiment begin(*>*) |
53536
69c943125fd3
do not expose internal flags to attribute name space;
wenzelm
parents:
53370
diff
changeset
|
307 |
bundle trace = [[simp_trace, linarith_trace, metis_trace, smt_trace]] |
47484 | 308 |
|
309 |
lemma "x = x" |
|
310 |
including trace by metis |
|
59905 | 311 |
(*<*)end(*>*) |
47484 | 312 |
|
313 |
||
58618 | 314 |
section \<open>Term definitions \label{sec:term-definitions}\<close> |
27040 | 315 |
|
58618 | 316 |
text \<open> |
27040 | 317 |
\begin{matharray}{rcll} |
61493 | 318 |
@{command_def "definition"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
319 |
@{attribute_def "defn"} & : & \<open>attribute\<close> \\ |
|
61494 | 320 |
@{command_def "print_defn_rules"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
61493 | 321 |
@{command_def "abbreviation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
61494 | 322 |
@{command_def "print_abbrevs"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
27040 | 323 |
\end{matharray} |
324 |
||
61664 | 325 |
Term definitions may either happen within the logic (as equational axioms of |
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
326 |
a certain form (see also \secref{sec:overloading}), or outside of it as |
57487 | 327 |
rewrite system on abstract syntax. The second form is called |
328 |
``abbreviation''. |
|
27040 | 329 |
|
69597 | 330 |
\<^rail>\<open> |
63180 | 331 |
@@{command definition} decl? definition |
27040 | 332 |
; |
63180 | 333 |
@@{command abbreviation} @{syntax mode}? decl? abbreviation |
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
334 |
; |
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
335 |
@@{command print_abbrevs} ('!'?) |
63180 | 336 |
; |
337 |
decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? @'where' |
|
338 |
; |
|
63182 | 339 |
definition: @{syntax thmdecl}? @{syntax prop} |
340 |
@{syntax spec_prems} @{syntax for_fixes} |
|
63180 | 341 |
; |
342 |
abbreviation: @{syntax prop} @{syntax for_fixes} |
|
69597 | 343 |
\<close> |
27040 | 344 |
|
61665 | 345 |
\<^descr> \<^theory_text>\<open>definition c where eq\<close> produces an internal definition \<open>c \<equiv> t\<close> according |
346 |
to the specification given as \<open>eq\<close>, which is then turned into a proven fact. |
|
347 |
The given proposition may deviate from internal meta-level equality |
|
348 |
according to the rewrite rules declared as @{attribute defn} by the |
|
61664 | 349 |
object-logic. This usually covers object-level equality \<open>x = y\<close> and |
61665 | 350 |
equivalence \<open>A \<longleftrightarrow> B\<close>. End-users normally need not change the @{attribute |
61664 | 351 |
defn} setup. |
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
352 |
|
61664 | 353 |
Definitions may be presented with explicit arguments on the LHS, as well as |
354 |
additional conditions, e.g.\ \<open>f x y = t\<close> instead of \<open>f \<equiv> \<lambda>x y. t\<close> and \<open>y \<noteq> 0 |
|
355 |
\<Longrightarrow> g x y = u\<close> instead of an unrestricted \<open>g \<equiv> \<lambda>x y. u\<close>. |
|
51585 | 356 |
|
61665 | 357 |
\<^descr> \<^theory_text>\<open>print_defn_rules\<close> prints the definitional rewrite rules declared via |
358 |
@{attribute defn} in the current context. |
|
51585 | 359 |
|
61665 | 360 |
\<^descr> \<^theory_text>\<open>abbreviation c where eq\<close> introduces a syntactic constant which is |
361 |
associated with a certain term according to the meta-level equality \<open>eq\<close>. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
362 |
|
61664 | 363 |
Abbreviations participate in the usual type-inference process, but are |
364 |
expanded before the logic ever sees them. Pretty printing of terms involves |
|
365 |
higher-order rewriting with rules stemming from reverted abbreviations. This |
|
366 |
needs some care to avoid overlapping or looping syntactic replacements! |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
367 |
|
61664 | 368 |
The optional \<open>mode\<close> specification restricts output to a particular print |
369 |
mode; using ``\<open>input\<close>'' here achieves the effect of one-way abbreviations. |
|
61665 | 370 |
The mode may also include an ``\<^theory_text>\<open>output\<close>'' qualifier that affects the |
371 |
concrete syntax declared for abbreviations, cf.\ \<^theory_text>\<open>syntax\<close> in |
|
27040 | 372 |
\secref{sec:syn-trans}. |
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
373 |
|
61665 | 374 |
\<^descr> \<^theory_text>\<open>print_abbrevs\<close> prints all constant abbreviations of the current context; |
375 |
the ``\<open>!\<close>'' option indicates extra verbosity. |
|
58618 | 376 |
\<close> |
27040 | 377 |
|
378 |
||
58618 | 379 |
section \<open>Axiomatizations \label{sec:axiomatizations}\<close> |
57487 | 380 |
|
58618 | 381 |
text \<open> |
57487 | 382 |
\begin{matharray}{rcll} |
61493 | 383 |
@{command_def "axiomatization"} & : & \<open>theory \<rightarrow> theory\<close> & (axiomatic!) \\ |
57487 | 384 |
\end{matharray} |
385 |
||
69597 | 386 |
\<^rail>\<open> |
63285 | 387 |
@@{command axiomatization} @{syntax vars}? (@'where' axiomatization)? |
57487 | 388 |
; |
63178 | 389 |
axiomatization: (@{syntax thmdecl} @{syntax prop} + @'and') |
63182 | 390 |
@{syntax spec_prems} @{syntax for_fixes} |
69597 | 391 |
\<close> |
57487 | 392 |
|
61665 | 393 |
\<^descr> \<^theory_text>\<open>axiomatization c\<^sub>1 \<dots> c\<^sub>m where \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n\<close> introduces several constants |
394 |
simultaneously and states axiomatic properties for these. The constants are |
|
395 |
marked as being specified once and for all, which prevents additional |
|
71567
9a29e883a934
tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents:
71166
diff
changeset
|
396 |
specifications for the same constants later on, but it is always possible to |
61665 | 397 |
emit axiomatizations without referring to particular constants. Note that |
398 |
lack of precise dependency tracking of axiomatizations may disrupt the |
|
399 |
well-formedness of an otherwise definitional theory. |
|
57487 | 400 |
|
401 |
Axiomatization is restricted to a global theory context: support for local |
|
402 |
theory targets \secref{sec:target} would introduce an extra dimension of |
|
403 |
uncertainty what the written specifications really are, and make it |
|
404 |
infeasible to argue why they are correct. |
|
405 |
||
406 |
Axiomatic specifications are required when declaring a new logical system |
|
407 |
within Isabelle/Pure, but in an application environment like Isabelle/HOL |
|
61664 | 408 |
the user normally stays within definitional mechanisms provided by the logic |
409 |
and its libraries. |
|
58618 | 410 |
\<close> |
57487 | 411 |
|
412 |
||
58618 | 413 |
section \<open>Generic declarations\<close> |
27040 | 414 |
|
58618 | 415 |
text \<open> |
47483 | 416 |
\begin{matharray}{rcl} |
61493 | 417 |
@{command_def "declaration"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
418 |
@{command_def "syntax_declaration"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
419 |
@{command_def "declare"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
47483 | 420 |
\end{matharray} |
421 |
||
61664 | 422 |
Arbitrary operations on the background context may be wrapped-up as generic |
423 |
declaration elements. Since the underlying concept of local theories may be |
|
424 |
subject to later re-interpretation, there is an additional dependency on a |
|
425 |
morphism that tells the difference of the original declaration context wrt.\ |
|
426 |
the application context encountered later on. A fact declaration is an |
|
427 |
important special case: it consists of a theorem which is applied to the |
|
428 |
context by means of an attribute. |
|
27040 | 429 |
|
69597 | 430 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
431 |
(@@{command declaration} | @@{command syntax_declaration}) |
59783
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
wenzelm
parents:
59781
diff
changeset
|
432 |
('(' @'pervasive' ')')? \<newline> @{syntax text} |
27040 | 433 |
; |
62969 | 434 |
@@{command declare} (@{syntax thms} + @'and') |
69597 | 435 |
\<close> |
27040 | 436 |
|
78085 | 437 |
\<^descr> \<^theory_text>\<open>declaration d\<close> adds the declaration function \<open>d\<close> of ML type \<^ML_type>\<open>Morphism.declaration\<close>, to the current local theory under construction. In later |
61665 | 438 |
application contexts, the function is transformed according to the morphisms |
439 |
being involved in the interpretation hierarchy. |
|
27040 | 440 |
|
61665 | 441 |
If the \<^theory_text>\<open>(pervasive)\<close> option is given, the corresponding declaration is |
61664 | 442 |
applied to all possible contexts involved, including the global background |
443 |
theory. |
|
33516 | 444 |
|
61665 | 445 |
\<^descr> \<^theory_text>\<open>syntax_declaration\<close> is similar to \<^theory_text>\<open>declaration\<close>, but is meant to affect |
446 |
only ``syntactic'' tools by convention (such as notation and type-checking |
|
447 |
information). |
|
40784 | 448 |
|
61665 | 449 |
\<^descr> \<^theory_text>\<open>declare thms\<close> declares theorems to the current local theory context. No |
450 |
theorem binding is involved here, unlike \<^theory_text>\<open>lemmas\<close> (cf.\ |
|
451 |
\secref{sec:theorems}), so \<^theory_text>\<open>declare\<close> only has the effect of applying |
|
452 |
attributes as included in the theorem specification. |
|
58618 | 453 |
\<close> |
27040 | 454 |
|
455 |
||
58618 | 456 |
section \<open>Locales \label{sec:locale}\<close> |
27040 | 457 |
|
58618 | 458 |
text \<open> |
54049 | 459 |
A locale is a functor that maps parameters (including implicit type |
61664 | 460 |
parameters) and a specification to a list of declarations. The syntax of |
461 |
locales is modeled after the Isar proof context commands (cf.\ |
|
462 |
\secref{sec:proof-context}). |
|
54049 | 463 |
|
61664 | 464 |
Locale hierarchies are supported by maintaining a graph of dependencies |
465 |
between locale instances in the global theory. Dependencies may be |
|
466 |
introduced through import (where a locale is defined as sublocale of the |
|
467 |
imported instances) or by proving that an existing locale is a sublocale of |
|
468 |
one or several locale instances. |
|
54049 | 469 |
|
470 |
A locale may be opened with the purpose of appending to its list of |
|
61664 | 471 |
declarations (cf.\ \secref{sec:target}). When opening a locale declarations |
472 |
from all dependencies are collected and are presented as a local theory. In |
|
473 |
this process, which is called \<^emph>\<open>roundup\<close>, redundant locale instances are |
|
474 |
omitted. A locale instance is redundant if it is subsumed by an instance |
|
475 |
encountered earlier. A more detailed description of this process is |
|
76987 | 476 |
available elsewhere \<^cite>\<open>Ballarin2014\<close>. |
58618 | 477 |
\<close> |
27040 | 478 |
|
479 |
||
58618 | 480 |
subsection \<open>Locale expressions \label{sec:locale-expr}\<close> |
33846 | 481 |
|
58618 | 482 |
text \<open> |
61606
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
483 |
A \<^emph>\<open>locale expression\<close> denotes a context composed of instances of existing |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
484 |
locales. The context consists of the declaration elements from the locale |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
485 |
instances. Redundant locale instances are omitted according to roundup. |
33846 | 486 |
|
69597 | 487 |
\<^rail>\<open> |
59785 | 488 |
@{syntax_def locale_expr}: (instance + '+') @{syntax for_fixes} |
33846 | 489 |
; |
67764 | 490 |
instance: (qualifier ':')? @{syntax name} (pos_insts | named_insts) \<newline> |
491 |
rewrites? |
|
33846 | 492 |
; |
61606
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
493 |
qualifier: @{syntax name} ('?')? |
33846 | 494 |
; |
42617 | 495 |
pos_insts: ('_' | @{syntax term})* |
33846 | 496 |
; |
42617 | 497 |
named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and') |
67740 | 498 |
; |
499 |
rewrites: @'rewrites' (@{syntax thmdecl}? @{syntax prop} + @'and') |
|
69597 | 500 |
\<close> |
33846 | 501 |
|
61606
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
502 |
A locale instance consists of a reference to a locale and either positional |
67740 | 503 |
or named parameter instantiations optionally followed by rewrites clauses. |
504 |
Identical instantiations (that is, those |
|
61606
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
505 |
that instantiate a parameter by itself) may be omitted. The notation ``\<open>_\<close>'' |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
506 |
enables to omit the instantiation for a parameter inside a positional |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
507 |
instantiation. |
33846 | 508 |
|
61606
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
509 |
Terms in instantiations are from the context the locale expressions is |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
510 |
declared in. Local names may be added to this context with the optional |
61665 | 511 |
\<^theory_text>\<open>for\<close> clause. This is useful for shadowing names bound in outer contexts, |
512 |
and for declaring syntax. In addition, syntax declarations from one instance |
|
513 |
are effective when parsing subsequent instances of the same expression. |
|
33846 | 514 |
|
61606
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
515 |
Instances have an optional qualifier which applies to names in declarations. |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
516 |
Names include local definitions and theorem names. If present, the qualifier |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
517 |
itself is either mandatory (default) or non-mandatory (when followed by |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
518 |
``\<^verbatim>\<open>?\<close>''). Non-mandatory means that the qualifier may be omitted on input. |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
519 |
Qualifiers only affect name spaces; they play no role in determining whether |
6d5213bd9709
uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents:
61566
diff
changeset
|
520 |
one locale instance subsumes another. |
67740 | 521 |
|
522 |
Rewrite clauses amend instances with equations that act as rewrite rules. |
|
523 |
This is particularly useful for changing concepts introduced through |
|
524 |
definitions. Rewrite clauses are available only in interpretation commands |
|
525 |
(see \secref{sec:locale-interpretation} below) and must be proved the user. |
|
58618 | 526 |
\<close> |
33846 | 527 |
|
528 |
||
58618 | 529 |
subsection \<open>Locale declarations\<close> |
27040 | 530 |
|
58618 | 531 |
text \<open> |
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
532 |
\begin{tabular}{rcl} |
61493 | 533 |
@{command_def "locale"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
534 |
@{command_def "experiment"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
|
535 |
@{command_def "print_locale"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
|
536 |
@{command_def "print_locales"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
|
537 |
@{command_def "locale_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
|
70608 | 538 |
\end{tabular} |
27040 | 539 |
|
73770
1419cb7f7f3e
clarified index, more like formal @{element_ref};
wenzelm
parents:
73757
diff
changeset
|
540 |
@{index_ref \<open>\<^theory_text>\<open>fixes\<close> (element)\<close>} |
1419cb7f7f3e
clarified index, more like formal @{element_ref};
wenzelm
parents:
73757
diff
changeset
|
541 |
@{index_ref \<open>\<^theory_text>\<open>constrains\<close> (element)\<close>} |
1419cb7f7f3e
clarified index, more like formal @{element_ref};
wenzelm
parents:
73757
diff
changeset
|
542 |
@{index_ref \<open>\<^theory_text>\<open>assumes\<close> (element)\<close>} |
1419cb7f7f3e
clarified index, more like formal @{element_ref};
wenzelm
parents:
73757
diff
changeset
|
543 |
@{index_ref \<open>\<^theory_text>\<open>defines\<close> (element)\<close>} |
1419cb7f7f3e
clarified index, more like formal @{element_ref};
wenzelm
parents:
73757
diff
changeset
|
544 |
@{index_ref \<open>\<^theory_text>\<open>notes\<close> (element)\<close>} |
69597 | 545 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
546 |
@@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'? |
27040 | 547 |
; |
59901 | 548 |
@@{command experiment} (@{syntax context_elem}*) @'begin' |
549 |
; |
|
62969 | 550 |
@@{command print_locale} '!'? @{syntax name} |
27040 | 551 |
; |
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
552 |
@@{command print_locales} ('!'?) |
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
553 |
; |
42617 | 554 |
@{syntax_def locale}: @{syntax context_elem}+ | |
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
555 |
@{syntax_ref "opening"} ('+' (@{syntax context_elem}+))? | |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
556 |
@{syntax locale_expr} @{syntax_ref "opening"}? ('+' (@{syntax context_elem}+))? |
27040 | 557 |
; |
42617 | 558 |
@{syntax_def context_elem}: |
63285 | 559 |
@'fixes' @{syntax vars} | |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
560 |
@'constrains' (@{syntax name} '::' @{syntax type} + @'and') | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
561 |
@'assumes' (@{syntax props} + @'and') | |
42705 | 562 |
@'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') | |
62969 | 563 |
@'notes' (@{syntax thmdef}? @{syntax thms} + @'and') |
69597 | 564 |
\<close> |
27040 | 565 |
|
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
566 |
\<^descr> \<^theory_text>\<open>locale loc = import opening bundles + body\<close> defines a new locale \<open>loc\<close> |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
567 |
as a context consisting of a certain view of existing locales (\<open>import\<close>) plus some |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
568 |
additional elements (\<open>body\<close>) with declaration \<open>bundles\<close> enriching the context |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
569 |
of the command itself. All \<open>import\<close>, \<open>bundles\<close> and \<open>body\<close> are optional; the |
61665 | 570 |
degenerate form \<^theory_text>\<open>locale loc\<close> defines an empty locale, which may still be |
571 |
useful to collect declarations of facts later on. Type-inference on locale |
|
572 |
expressions automatically takes care of the most general typing that the |
|
573 |
combined context elements may acquire. |
|
27040 | 574 |
|
71567
9a29e883a934
tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents:
71166
diff
changeset
|
575 |
The \<open>import\<close> consists of a locale expression; see \secref{sec:locale-expr} |
61665 | 576 |
above. Its \<^theory_text>\<open>for\<close> clause defines the parameters of \<open>import\<close>. These are |
577 |
parameters of the defined locale. Locale parameters whose instantiation is |
|
578 |
omitted automatically extend the (possibly empty) \<^theory_text>\<open>for\<close> clause: they are |
|
579 |
inserted at its beginning. This means that these parameters may be referred |
|
580 |
to from within the expression and also in the subsequent context elements |
|
581 |
and provides a notational convenience for the inheritance of parameters in |
|
582 |
locale declarations. |
|
27040 | 583 |
|
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
584 |
Declarations from \<open>bundles\<close>, see \secref{sec:bundle}, are effective in the |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
585 |
entire command including a subsequent \<^theory_text>\<open>begin\<close> / \<^theory_text>\<open>end\<close> block, but they do |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
586 |
not contribute to the declarations stored in the locale. |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
587 |
|
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
588 |
The \<open>body\<close> consists of context elements: |
27040 | 589 |
|
61664 | 590 |
\<^descr> @{element "fixes"}~\<open>x :: \<tau> (mx)\<close> declares a local parameter of type \<open>\<tau>\<close> |
591 |
and mixfix annotation \<open>mx\<close> (both are optional). The special syntax |
|
592 |
declaration ``\<open>(\<close>@{keyword_ref "structure"}\<open>)\<close>'' means that \<open>x\<close> may be |
|
593 |
referenced implicitly in this context. |
|
27040 | 594 |
|
61664 | 595 |
\<^descr> @{element "constrains"}~\<open>x :: \<tau>\<close> introduces a type constraint \<open>\<tau>\<close> on the |
596 |
local parameter \<open>x\<close>. This element is deprecated. The type constraint |
|
61665 | 597 |
should be introduced in the \<^theory_text>\<open>for\<close> clause or the relevant @{element |
598 |
"fixes"} element. |
|
27040 | 599 |
|
61664 | 600 |
\<^descr> @{element "assumes"}~\<open>a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n\<close> introduces local premises, similar |
61665 | 601 |
to \<^theory_text>\<open>assume\<close> within a proof (cf.\ \secref{sec:proof-context}). |
27040 | 602 |
|
61664 | 603 |
\<^descr> @{element "defines"}~\<open>a: x \<equiv> t\<close> defines a previously declared parameter. |
63039 | 604 |
This is similar to \<^theory_text>\<open>define\<close> within a proof (cf.\ |
605 |
\secref{sec:proof-context}), but @{element "defines"} is restricted to |
|
606 |
Pure equalities and the defined variable needs to be declared beforehand |
|
607 |
via @{element "fixes"}. The left-hand side of the equation may have |
|
608 |
additional arguments, e.g.\ ``@{element "defines"}~\<open>f x\<^sub>1 \<dots> x\<^sub>n \<equiv> t\<close>'', |
|
609 |
which need to be free in the context. |
|
27040 | 610 |
|
61664 | 611 |
\<^descr> @{element "notes"}~\<open>a = b\<^sub>1 \<dots> b\<^sub>n\<close> reconsiders facts within a local |
612 |
context. Most notably, this may include arbitrary declarations in any |
|
613 |
attribute specifications included here, e.g.\ a local @{attribute simp} |
|
614 |
rule. |
|
27040 | 615 |
|
61664 | 616 |
Both @{element "assumes"} and @{element "defines"} elements contribute to |
617 |
the locale specification. When defining an operation derived from the |
|
61665 | 618 |
parameters, \<^theory_text>\<open>definition\<close> (\secref{sec:term-definitions}) is usually more |
619 |
appropriate. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
620 |
|
61665 | 621 |
Note that ``\<^theory_text>\<open>(is p\<^sub>1 \<dots> p\<^sub>n)\<close>'' patterns given in the syntax of @{element |
61664 | 622 |
"assumes"} and @{element "defines"} above are illegal in locale definitions. |
623 |
In the long goal format of \secref{sec:goals}, term bindings may be included |
|
624 |
as expected, though. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
625 |
|
61421 | 626 |
\<^medskip> |
61664 | 627 |
Locale specifications are ``closed up'' by turning the given text into a |
628 |
predicate definition \<open>loc_axioms\<close> and deriving the original assumptions as |
|
629 |
local lemmas (modulo local definitions). The predicate statement covers only |
|
630 |
the newly specified assumptions, omitting the content of included locale |
|
631 |
expressions. The full cumulative view is only provided on export, involving |
|
632 |
another predicate \<open>loc\<close> that refers to the complete specification text. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
633 |
|
61664 | 634 |
In any case, the predicate arguments are those locale parameters that |
635 |
actually occur in the respective piece of text. Also these predicates |
|
636 |
operate at the meta-level in theory, but the locale packages attempts to |
|
637 |
internalize statements according to the object-logic setup (e.g.\ replacing |
|
638 |
\<open>\<And>\<close> by \<open>\<forall>\<close>, and \<open>\<Longrightarrow>\<close> by \<open>\<longrightarrow>\<close> in HOL; see also \secref{sec:object-logic}). |
|
639 |
Separate introduction rules \<open>loc_axioms.intro\<close> and \<open>loc.intro\<close> are provided |
|
640 |
as well. |
|
59901 | 641 |
|
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
642 |
\<^descr> \<^theory_text>\<open>experiment body begin\<close> opens an anonymous locale context with private |
61665 | 643 |
naming policy. Specifications in its body are inaccessible from outside. |
644 |
This is useful to perform experiments, without polluting the name space. |
|
59901 | 645 |
|
69018 | 646 |
\<^descr> \<^theory_text>\<open>print_locale "locale"\<close> prints the contents of the named locale. The |
61665 | 647 |
command omits @{element "notes"} elements by default. Use \<^theory_text>\<open>print_locale!\<close> |
648 |
to have them included. |
|
27040 | 649 |
|
61665 | 650 |
\<^descr> \<^theory_text>\<open>print_locales\<close> prints the names of all locales of the current theory; |
651 |
the ``\<open>!\<close>'' option indicates extra verbosity. |
|
27040 | 652 |
|
61665 | 653 |
\<^descr> \<^theory_text>\<open>locale_deps\<close> visualizes all locales and their relations as a Hasse |
654 |
diagram. This includes locales defined as type classes (\secref{sec:class}). |
|
58618 | 655 |
\<close> |
27040 | 656 |
|
657 |
||
67740 | 658 |
subsection \<open>Locale interpretation \label{sec:locale-interpretation}\<close> |
27040 | 659 |
|
58618 | 660 |
text \<open> |
27040 | 661 |
\begin{matharray}{rcl} |
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
662 |
@{command "interpretation"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close> \\ |
61493 | 663 |
@{command_def "interpret"} & : & \<open>proof(state) | proof(chain) \<rightarrow> proof(prove)\<close> \\ |
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
664 |
@{command_def "global_interpretation"} & : & \<open>theory | local_theory \<rightarrow> proof(prove)\<close> \\ |
61493 | 665 |
@{command_def "sublocale"} & : & \<open>theory | local_theory \<rightarrow> proof(prove)\<close> \\ |
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
666 |
@{command_def "print_interps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
667 |
@{method_def intro_locales} & : & \<open>method\<close> \\ |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
668 |
@{method_def unfold_locales} & : & \<open>method\<close> \\ |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
669 |
@{attribute_def trace_locales} & : & \mbox{\<open>attribute\<close> \quad default \<open>false\<close>} \\ |
27040 | 670 |
\end{matharray} |
671 |
||
61664 | 672 |
Locales may be instantiated, and the resulting instantiated declarations |
673 |
added to the current context. This requires proof (of the instantiated |
|
61706 | 674 |
specification) and is called \<^emph>\<open>locale interpretation\<close>. Interpretation is |
675 |
possible within arbitrary local theories (\<^theory_text>\<open>interpretation\<close>), within proof |
|
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
676 |
bodies (\<^theory_text>\<open>interpret\<close>), into global theories (\<^theory_text>\<open>global_interpretation\<close>) and |
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
677 |
into locales (\<^theory_text>\<open>sublocale\<close>). |
47483 | 678 |
|
69597 | 679 |
\<^rail>\<open> |
67740 | 680 |
@@{command interpretation} @{syntax locale_expr} |
27040 | 681 |
; |
67740 | 682 |
@@{command interpret} @{syntax locale_expr} |
27040 | 683 |
; |
67764 | 684 |
@@{command global_interpretation} @{syntax locale_expr} definitions? |
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
685 |
; |
62969 | 686 |
@@{command sublocale} (@{syntax name} ('<' | '\<subseteq>'))? @{syntax locale_expr} \<newline> |
67764 | 687 |
definitions? |
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
688 |
; |
62969 | 689 |
@@{command print_interps} @{syntax name} |
41434
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
ballarin
parents:
41249
diff
changeset
|
690 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
691 |
|
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
692 |
definitions: @'defines' (@{syntax thmdecl}? @{syntax name} \<newline> |
70058 | 693 |
@{syntax mixfix}? '=' @{syntax term} + @'and'); |
69597 | 694 |
\<close> |
27040 | 695 |
|
61706 | 696 |
The core of each interpretation command is a locale expression \<open>expr\<close>; the |
697 |
command generates proof obligations for the instantiated specifications. |
|
698 |
Once these are discharged by the user, instantiated declarations (in |
|
699 |
particular, facts) are added to the context in a post-processing phase, in a |
|
700 |
manner specific to each command. |
|
53368
92b9965f479d
Clarifies documentation of interpretation in local theories.
ballarin
parents:
53015
diff
changeset
|
701 |
|
61706 | 702 |
Interpretation commands are aware of interpretations that are already |
703 |
active: post-processing is achieved through a variant of roundup that takes |
|
61664 | 704 |
interpretations of the current global or local theory into account. In order |
705 |
to simplify the proof obligations according to existing interpretations use |
|
706 |
methods @{method intro_locales} or @{method unfold_locales}. |
|
53368
92b9965f479d
Clarifies documentation of interpretation in local theories.
ballarin
parents:
53015
diff
changeset
|
707 |
|
67764 | 708 |
Rewrites clauses \<^theory_text>\<open>rewrites eqns\<close> occur within expressions. They amend the |
709 |
morphism through which a locale instance is interpreted with rewrite rules, |
|
710 |
also called rewrite morphisms. This is particularly useful for interpreting |
|
711 |
concepts introduced through definitions. The equations must be proved the |
|
712 |
user. To enable syntax of the instantiated locale within the equation, while |
|
713 |
reading a locale expression, equations of a locale instance are read in a |
|
714 |
temporary context where the instance is already activated. If activation |
|
715 |
fails, typically due to duplicate constant declarations, processing falls |
|
716 |
back to reading the equation first. |
|
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
717 |
|
61706 | 718 |
Given definitions \<open>defs\<close> produce corresponding definitions in the local |
67740 | 719 |
theory's underlying target \<^emph>\<open>and\<close> amend the morphism with rewrite rules |
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
720 |
stemming from the symmetric of those definitions. Hence these need not be |
61706 | 721 |
proved explicitly the user. Such rewrite definitions are a even more useful |
722 |
device for interpreting concepts introduced through definitions, but they |
|
723 |
are only supported for interpretation commands operating in a local theory |
|
61823 | 724 |
whose implementing target actually supports this. Note that despite |
725 |
the suggestive \<^theory_text>\<open>and\<close> connective, \<open>defs\<close> |
|
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
726 |
are processed sequentially without mutual recursion. |
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
727 |
|
67740 | 728 |
\<^descr> \<^theory_text>\<open>interpretation expr\<close> interprets \<open>expr\<close> into a local theory |
61706 | 729 |
such that its lifetime is limited to the current context block (e.g. a |
730 |
locale or unnamed context). At the closing @{command end} of the block the |
|
731 |
interpretation and its declarations disappear. Hence facts based on |
|
732 |
interpretation can be established without creating permanent links to the |
|
733 |
interpreted locale instances, as would be the case with @{command |
|
734 |
sublocale}. |
|
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
735 |
|
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
736 |
When used on the level of a global theory, there is no end of a current |
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
737 |
context block, hence \<^theory_text>\<open>interpretation\<close> behaves identically to |
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
738 |
\<^theory_text>\<open>global_interpretation\<close> then. |
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
739 |
|
67740 | 740 |
\<^descr> \<^theory_text>\<open>interpret expr\<close> interprets \<open>expr\<close> into a proof context: |
61706 | 741 |
the interpretation and its declarations disappear when closing the current |
742 |
proof block. Note that for \<^theory_text>\<open>interpret\<close> the \<open>eqns\<close> should be explicitly |
|
743 |
universally quantified. |
|
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
744 |
|
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
745 |
\<^descr> \<^theory_text>\<open>global_interpretation expr defines defs\<close> interprets \<open>expr\<close> |
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
746 |
into a global theory. |
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
747 |
|
53369
b4bcac7d065b
Clarifies that interpretation does not only apply to facts, but to declaratoins in general.
ballarin
parents:
53368
diff
changeset
|
748 |
When adding declarations to locales, interpreted versions of these |
61664 | 749 |
declarations are added to the global theory for all interpretations in the |
61891
76189756ff65
documentation on last state of the art concerning interpretation
haftmann
parents:
61853
diff
changeset
|
750 |
global theory as well. That is, interpretations into global theories |
61664 | 751 |
dynamically participate in any declarations added to locales. |
54049 | 752 |
|
61664 | 753 |
Free variables in the interpreted expression are allowed. They are turned |
754 |
into schematic variables in the generated declarations. In order to use a |
|
755 |
free variable whose name is already bound in the context --- for example, |
|
61665 | 756 |
because a constant of that name exists --- add it to the \<^theory_text>\<open>for\<close> clause. |
48824
45d0e40b07af
Clarification: free variables allowed in interpreted locale instances.
ballarin
parents:
47484
diff
changeset
|
757 |
|
73846 | 758 |
When used in a nested target, resulting declarations are propagated |
759 |
through the whole target stack. |
|
760 |
||
67764 | 761 |
\<^descr> \<^theory_text>\<open>sublocale name \<subseteq> expr defines defs\<close> interprets \<open>expr\<close> |
61706 | 762 |
into the locale \<open>name\<close>. A proof that the specification of \<open>name\<close> implies the |
763 |
specification of \<open>expr\<close> is required. As in the localized version of the |
|
764 |
theorem command, the proof is in the context of \<open>name\<close>. After the proof |
|
765 |
obligation has been discharged, the locale hierarchy is changed as if \<open>name\<close> |
|
766 |
imported \<open>expr\<close> (hence the name \<^theory_text>\<open>sublocale\<close>). When the context of \<open>name\<close> is |
|
767 |
subsequently entered, traversing the locale hierarchy will involve the |
|
768 |
locale instances of \<open>expr\<close>, and their declarations will be added to the |
|
769 |
context. This makes \<^theory_text>\<open>sublocale\<close> dynamic: extensions of a locale that is |
|
770 |
instantiated in \<open>expr\<close> may take place after the \<^theory_text>\<open>sublocale\<close> declaration and |
|
771 |
still become available in the context. Circular \<^theory_text>\<open>sublocale\<close> declarations |
|
772 |
are allowed as long as they do not lead to infinite chains. |
|
41434
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
ballarin
parents:
41249
diff
changeset
|
773 |
|
61664 | 774 |
If interpretations of \<open>name\<close> exist in the current global theory, the command |
775 |
adds interpretations for \<open>expr\<close> as well, with the same qualifier, although |
|
776 |
only for fragments of \<open>expr\<close> that are not interpreted in the theory already. |
|
41434
710cdb9e0d17
Documentation for 'interpret' and 'sublocale' with mixins.
ballarin
parents:
41249
diff
changeset
|
777 |
|
67764 | 778 |
Rewrites clauses in the expression or rewrite definitions \<open>defs\<close> can help break |
67740 | 779 |
infinite chains induced by circular \<^theory_text>\<open>sublocale\<close> declarations. |
27040 | 780 |
|
61665 | 781 |
In a named context block the \<^theory_text>\<open>sublocale\<close> command may also be used, but the |
782 |
locale argument must be omitted. The command then refers to the locale (or |
|
783 |
class) target of the context block. |
|
51747 | 784 |
|
71166 | 785 |
\<^descr> \<^theory_text>\<open>print_interps name\<close> lists all interpretations of locale \<open>name\<close> in the |
61665 | 786 |
current theory or proof context, including those due to a combination of an |
787 |
\<^theory_text>\<open>interpretation\<close> or \<^theory_text>\<open>interpret\<close> and one or several \<^theory_text>\<open>sublocale\<close> |
|
788 |
declarations. |
|
33867 | 789 |
|
72804
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
790 |
\<^descr> @{method intro_locales} and @{method unfold_locales} repeatedly expand all |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
791 |
introduction rules of locale predicates of the theory. While @{method |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
792 |
intro_locales} only applies the \<open>loc.intro\<close> introduction rules and therefore |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
793 |
does not descend to assumptions, @{method unfold_locales} is more aggressive |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
794 |
and applies \<open>loc_axioms.intro\<close> as well. Both methods are aware of locale |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
795 |
specifications entailed by the context, both from target statements, and |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
796 |
from interpretations (see below). New goals that are entailed by the current |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
797 |
context are discharged automatically. |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
798 |
|
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
799 |
While @{method unfold_locales} is part of the default method for \<^theory_text>\<open>proof\<close> |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
800 |
and often invoked ``behind the scenes'', @{method intro_locales} helps |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
801 |
understand which proof obligations originated from which locale instances. |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
802 |
The latter method is useful while developing proofs but rare in finished |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
803 |
developments. |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
804 |
|
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
805 |
\<^descr> @{attribute trace_locales}, when set to \<open>true\<close>, prints the locale |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
806 |
instances activated during roundup. Use this when locale commands yield |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
807 |
obscure errors or for understanding local theories created by complex locale |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
808 |
hierarchies. |
d7393c35aa5d
Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents:
72739
diff
changeset
|
809 |
|
27040 | 810 |
\begin{warn} |
61664 | 811 |
If a global theory inherits declarations (body elements) for a locale from |
812 |
one parent and an interpretation of that locale from another parent, the |
|
813 |
interpretation will not be applied to the declarations. |
|
54049 | 814 |
\end{warn} |
815 |
||
816 |
\begin{warn} |
|
61664 | 817 |
Since attributes are applied to interpreted theorems, interpretation may |
818 |
modify the context of common proof tools, e.g.\ the Simplifier or |
|
61665 | 819 |
Classical Reasoner. As the behaviour of such tools is \<^emph>\<open>not\<close> stable under |
61664 | 820 |
interpretation morphisms, manual declarations might have to be added to |
821 |
the target context of the interpretation to revert such declarations. |
|
27040 | 822 |
\end{warn} |
823 |
||
824 |
\begin{warn} |
|
51747 | 825 |
An interpretation in a local theory or proof context may subsume previous |
61664 | 826 |
interpretations. This happens if the same specification fragment is |
827 |
interpreted twice and the instantiation of the second interpretation is |
|
828 |
more general than the interpretation of the first. The locale package does |
|
829 |
not attempt to remove subsumed interpretations. |
|
27040 | 830 |
\end{warn} |
831 |
||
61673
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
832 |
\begin{warn} |
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
833 |
While \<^theory_text>\<open>interpretation (in c) \<dots>\<close> is admissible, it is not useful since |
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
834 |
its result is discarded immediately. |
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents:
61671
diff
changeset
|
835 |
\end{warn} |
59003
16d92d37a8a1
documentation stubs about permanent_interpretation
haftmann
parents:
58724
diff
changeset
|
836 |
\<close> |
16d92d37a8a1
documentation stubs about permanent_interpretation
haftmann
parents:
58724
diff
changeset
|
837 |
|
16d92d37a8a1
documentation stubs about permanent_interpretation
haftmann
parents:
58724
diff
changeset
|
838 |
|
58618 | 839 |
section \<open>Classes \label{sec:class}\<close> |
27040 | 840 |
|
58618 | 841 |
text \<open> |
27040 | 842 |
\begin{matharray}{rcl} |
61493 | 843 |
@{command_def "class"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
844 |
@{command_def "instantiation"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
|
845 |
@{command_def "instance"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
846 |
@{command "instance"} & : & \<open>theory \<rightarrow> proof(prove)\<close> \\ |
|
847 |
@{command_def "subclass"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
848 |
@{command_def "print_classes"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
|
849 |
@{command_def "class_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
|
850 |
@{method_def intro_classes} & : & \<open>method\<close> \\ |
|
27040 | 851 |
\end{matharray} |
852 |
||
61664 | 853 |
A class is a particular locale with \<^emph>\<open>exactly one\<close> type variable \<open>\<alpha>\<close>. Beyond |
854 |
the underlying locale, a corresponding type class is established which is |
|
76987 | 855 |
interpreted logically as axiomatic type class \<^cite>\<open>"Wenzel:1997:TPHOL"\<close> |
61664 | 856 |
whose logical content are the assumptions of the locale. Thus, classes |
857 |
provide the full generality of locales combined with the commodity of type |
|
76987 | 858 |
classes (notably type-inference). See \<^cite>\<open>"isabelle-classes"\<close> for a short |
47483 | 859 |
tutorial. |
860 |
||
69597 | 861 |
\<^rail>\<open> |
42704 | 862 |
@@{command class} class_spec @'begin'? |
863 |
; |
|
864 |
class_spec: @{syntax name} '=' |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
865 |
((@{syntax name} @{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) | |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
866 |
@{syntax name} @{syntax_ref "opening"}? | |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
867 |
@{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) |
27040 | 868 |
; |
62969 | 869 |
@@{command instantiation} (@{syntax name} + @'and') '::' @{syntax arity} @'begin' |
27040 | 870 |
; |
62969 | 871 |
@@{command instance} (() | (@{syntax name} + @'and') '::' @{syntax arity} | |
872 |
@{syntax name} ('<' | '\<subseteq>') @{syntax name} ) |
|
31681 | 873 |
; |
62969 | 874 |
@@{command subclass} @{syntax name} |
58202 | 875 |
; |
60091 | 876 |
@@{command class_deps} (class_bounds class_bounds?)? |
60090 | 877 |
; |
60091 | 878 |
class_bounds: @{syntax sort} | '(' (@{syntax sort} + @'|') ')' |
69597 | 879 |
\<close> |
27040 | 880 |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
881 |
\<^descr> \<^theory_text>\<open>class c = superclasses bundles + body\<close> defines a new class \<open>c\<close>, inheriting from |
61665 | 882 |
\<open>superclasses\<close>. This introduces a locale \<open>c\<close> with import of all locales |
883 |
\<open>superclasses\<close>. |
|
27040 | 884 |
|
61664 | 885 |
Any @{element "fixes"} in \<open>body\<close> are lifted to the global theory level |
886 |
(\<^emph>\<open>class operations\<close> \<open>f\<^sub>1, \<dots>, f\<^sub>n\<close> of class \<open>c\<close>), mapping the local type |
|
887 |
parameter \<open>\<alpha>\<close> to a schematic type variable \<open>?\<alpha> :: c\<close>. |
|
27040 | 888 |
|
61664 | 889 |
Likewise, @{element "assumes"} in \<open>body\<close> are also lifted, mapping each local |
890 |
parameter \<open>f :: \<tau>[\<alpha>]\<close> to its corresponding global constant \<open>f :: \<tau>[?\<alpha> :: |
|
891 |
c]\<close>. The corresponding introduction rule is provided as |
|
892 |
\<open>c_class_axioms.intro\<close>. This rule should be rarely needed directly --- the |
|
893 |
@{method intro_classes} method takes care of the details of class membership |
|
894 |
proofs. |
|
27040 | 895 |
|
72739
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
896 |
Optionally given \<open>bundles\<close> take effect in the surface context within the |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
897 |
\<open>body\<close> and the potentially following \<^theory_text>\<open>begin\<close> / \<^theory_text>\<open>end\<close> block. |
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
haftmann
parents:
71926
diff
changeset
|
898 |
|
61665 | 899 |
\<^descr> \<^theory_text>\<open>instantiation t :: (s\<^sub>1, \<dots>, s\<^sub>n)s begin\<close> opens a target (cf.\ |
900 |
\secref{sec:target}) which allows to specify class operations \<open>f\<^sub>1, \<dots>, f\<^sub>n\<close> |
|
901 |
corresponding to sort \<open>s\<close> at the particular type instance \<open>(\<alpha>\<^sub>1 :: s\<^sub>1, \<dots>, |
|
902 |
\<alpha>\<^sub>n :: s\<^sub>n) t\<close>. A plain \<^theory_text>\<open>instance\<close> command in the target body poses a goal |
|
903 |
stating these type arities. The target is concluded by an @{command_ref |
|
904 |
(local) "end"} command. |
|
27040 | 905 |
|
61664 | 906 |
Note that a list of simultaneous type constructors may be given; this |
907 |
corresponds nicely to mutually recursive type definitions, e.g.\ in |
|
908 |
Isabelle/HOL. |
|
27040 | 909 |
|
61665 | 910 |
\<^descr> \<^theory_text>\<open>instance\<close> in an instantiation target body sets up a goal stating the |
911 |
type arities claimed at the opening \<^theory_text>\<open>instantiation\<close>. The proof would |
|
912 |
usually proceed by @{method intro_classes}, and then establish the |
|
913 |
characteristic theorems of the type classes involved. After finishing the |
|
914 |
proof, the background theory will be augmented by the proven type arities. |
|
31681 | 915 |
|
61665 | 916 |
On the theory level, \<^theory_text>\<open>instance t :: (s\<^sub>1, \<dots>, s\<^sub>n)s\<close> provides a convenient |
917 |
way to instantiate a type class with no need to specify operations: one can |
|
918 |
continue with the instantiation proof immediately. |
|
27040 | 919 |
|
61665 | 920 |
\<^descr> \<^theory_text>\<open>subclass c\<close> in a class context for class \<open>d\<close> sets up a goal stating that |
921 |
class \<open>c\<close> is logically contained in class \<open>d\<close>. After finishing the proof, |
|
922 |
class \<open>d\<close> is proven to be subclass \<open>c\<close> and the locale \<open>c\<close> is interpreted |
|
923 |
into \<open>d\<close> simultaneously. |
|
31681 | 924 |
|
61664 | 925 |
A weakened form of this is available through a further variant of @{command |
61665 | 926 |
instance}: \<^theory_text>\<open>instance c\<^sub>1 \<subseteq> c\<^sub>2\<close> opens a proof that class \<open>c\<^sub>2\<close> implies |
927 |
\<open>c\<^sub>1\<close> without reference to the underlying locales; this is useful if the |
|
928 |
properties to prove the logical connection are not sufficient on the locale |
|
929 |
level but on the theory level. |
|
61664 | 930 |
|
61665 | 931 |
\<^descr> \<^theory_text>\<open>print_classes\<close> prints all classes in the current theory. |
27040 | 932 |
|
61665 | 933 |
\<^descr> \<^theory_text>\<open>class_deps\<close> visualizes classes and their subclass relations as a |
934 |
directed acyclic graph. By default, all classes from the current theory |
|
61664 | 935 |
context are show. This may be restricted by optional bounds as follows: |
61665 | 936 |
\<^theory_text>\<open>class_deps upper\<close> or \<^theory_text>\<open>class_deps upper lower\<close>. A class is visualized, iff |
937 |
it is a subclass of some sort from \<open>upper\<close> and a superclass of some sort |
|
938 |
from \<open>lower\<close>. |
|
29706 | 939 |
|
61664 | 940 |
\<^descr> @{method intro_classes} repeatedly expands all class introduction rules of |
941 |
this theory. Note that this method usually needs not be named explicitly, as |
|
61665 | 942 |
it is already included in the default proof step (e.g.\ of \<^theory_text>\<open>proof\<close>). In |
943 |
particular, instantiation of trivial (syntactic) classes may be performed by |
|
944 |
a single ``\<^theory_text>\<open>..\<close>'' proof step. |
|
58618 | 945 |
\<close> |
26870 | 946 |
|
27040 | 947 |
|
58618 | 948 |
subsection \<open>The class target\<close> |
27040 | 949 |
|
58618 | 950 |
text \<open> |
27040 | 951 |
%FIXME check |
952 |
||
61664 | 953 |
A named context may refer to a locale (cf.\ \secref{sec:target}). If this |
954 |
locale is also a class \<open>c\<close>, apart from the common locale target behaviour |
|
955 |
the following happens. |
|
27040 | 956 |
|
61664 | 957 |
\<^item> Local constant declarations \<open>g[\<alpha>]\<close> referring to the local type parameter |
958 |
\<open>\<alpha>\<close> and local parameters \<open>f[\<alpha>]\<close> are accompanied by theory-level constants |
|
959 |
\<open>g[?\<alpha> :: c]\<close> referring to theory-level class operations \<open>f[?\<alpha> :: c]\<close>. |
|
27040 | 960 |
|
61664 | 961 |
\<^item> Local theorem bindings are lifted as are assumptions. |
27040 | 962 |
|
61664 | 963 |
\<^item> Local syntax refers to local operations \<open>g[\<alpha>]\<close> and global operations |
964 |
\<open>g[?\<alpha> :: c]\<close> uniformly. Type inference resolves ambiguities. In rare |
|
965 |
cases, manual type annotations are needed. |
|
58618 | 966 |
\<close> |
27040 | 967 |
|
968 |
||
58618 | 969 |
subsection \<open>Co-regularity of type classes and arities\<close> |
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
970 |
|
61664 | 971 |
text \<open> |
972 |
The class relation together with the collection of type-constructor arities |
|
973 |
must obey the principle of \<^emph>\<open>co-regularity\<close> as defined below. |
|
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
974 |
|
61421 | 975 |
\<^medskip> |
61664 | 976 |
For the subsequent formulation of co-regularity we assume that the class |
977 |
relation is closed by transitivity and reflexivity. Moreover the collection |
|
978 |
of arities \<open>t :: (\<^vec>s)c\<close> is completed such that \<open>t :: (\<^vec>s)c\<close> and |
|
979 |
\<open>c \<subseteq> c'\<close> implies \<open>t :: (\<^vec>s)c'\<close> for all such declarations. |
|
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
980 |
|
61664 | 981 |
Treating sorts as finite sets of classes (meaning the intersection), the |
982 |
class relation \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close> is extended to sorts as follows: |
|
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
983 |
\[ |
61493 | 984 |
\<open>s\<^sub>1 \<subseteq> s\<^sub>2 \<equiv> \<forall>c\<^sub>2 \<in> s\<^sub>2. \<exists>c\<^sub>1 \<in> s\<^sub>1. c\<^sub>1 \<subseteq> c\<^sub>2\<close> |
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
985 |
\] |
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
986 |
|
61664 | 987 |
This relation on sorts is further extended to tuples of sorts (of the same |
988 |
length) in the component-wise way. |
|
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
989 |
|
61421 | 990 |
\<^medskip> |
61664 | 991 |
Co-regularity of the class relation together with the arities relation |
992 |
means: |
|
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
993 |
\[ |
61493 | 994 |
\<open>t :: (\<^vec>s\<^sub>1)c\<^sub>1 \<Longrightarrow> t :: (\<^vec>s\<^sub>2)c\<^sub>2 \<Longrightarrow> c\<^sub>1 \<subseteq> c\<^sub>2 \<Longrightarrow> \<^vec>s\<^sub>1 \<subseteq> \<^vec>s\<^sub>2\<close> |
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
995 |
\] |
61664 | 996 |
for all such arities. In other words, whenever the result classes of some |
997 |
type-constructor arities are related, then the argument sorts need to be |
|
998 |
related in the same way. |
|
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
999 |
|
61421 | 1000 |
\<^medskip> |
61664 | 1001 |
Co-regularity is a very fundamental property of the order-sorted algebra of |
71567
9a29e883a934
tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents:
71166
diff
changeset
|
1002 |
types. For example, it entails principal types and most general unifiers, |
76987 | 1003 |
e.g.\ see \<^cite>\<open>"nipkow-prehofer"\<close>. |
58618 | 1004 |
\<close> |
37768
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
1005 |
|
e86221f3bac7
moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents:
37112
diff
changeset
|
1006 |
|
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1007 |
section \<open>Overloaded constant definitions \label{sec:overloading}\<close> |
27040 | 1008 |
|
58618 | 1009 |
text \<open> |
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1010 |
Definitions essentially express abbreviations within the logic. The simplest |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1011 |
form of a definition is \<open>c :: \<sigma> \<equiv> t\<close>, where \<open>c\<close> is a new constant and \<open>t\<close> is |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1012 |
a closed term that does not mention \<open>c\<close>. Moreover, so-called \<^emph>\<open>hidden |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1013 |
polymorphism\<close> is excluded: all type variables in \<open>t\<close> need to occur in its |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1014 |
type \<open>\<sigma>\<close>. |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1015 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1016 |
\<^emph>\<open>Overloading\<close> means that a constant being declared as \<open>c :: \<alpha> decl\<close> may be |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1017 |
defined separately on type instances \<open>c :: (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n)\<kappa> decl\<close> for each |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1018 |
type constructor \<open>\<kappa>\<close>. At most occasions overloading will be used in a |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1019 |
Haskell-like fashion together with type classes by means of \<^theory_text>\<open>instantiation\<close> |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1020 |
(see \secref{sec:class}). Sometimes low-level overloading is desirable; this |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1021 |
is supported by \<^theory_text>\<open>consts\<close> and \<^theory_text>\<open>overloading\<close> explained below. |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1022 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1023 |
The right-hand side of overloaded definitions may mention overloaded |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1024 |
constants recursively at type instances corresponding to the immediate |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1025 |
argument types \<open>\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n\<close>. Incomplete specification patterns impose |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1026 |
global constraints on all occurrences. E.g.\ \<open>d :: \<alpha> \<times> \<alpha>\<close> on the left-hand |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1027 |
side means that all corresponding occurrences on some right-hand side need |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1028 |
to be an instance of this, and general \<open>d :: \<alpha> \<times> \<beta>\<close> will be disallowed. Full |
76987 | 1029 |
details are given by Kun\v{c}ar \<^cite>\<open>"Kuncar:2015"\<close>. |
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1030 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1031 |
\<^medskip> |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1032 |
The \<^theory_text>\<open>consts\<close> command and the \<^theory_text>\<open>overloading\<close> target provide a convenient |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1033 |
interface for end-users. Regular specification elements such as @{command |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1034 |
definition}, @{command inductive}, @{command function} may be used in the |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1035 |
body. It is also possible to use \<^theory_text>\<open>consts c :: \<sigma>\<close> with later \<^theory_text>\<open>overloading c |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1036 |
\<equiv> c :: \<sigma>\<close> to keep the declaration and definition of a constant separate. |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1037 |
|
47483 | 1038 |
\begin{matharray}{rcl} |
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1039 |
@{command_def "consts"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
61493 | 1040 |
@{command_def "overloading"} & : & \<open>theory \<rightarrow> local_theory\<close> \\ |
47483 | 1041 |
\end{matharray} |
1042 |
||
69597 | 1043 |
\<^rail>\<open> |
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1044 |
@@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +) |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1045 |
; |
42704 | 1046 |
@@{command overloading} ( spec + ) @'begin' |
1047 |
; |
|
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1048 |
spec: @{syntax name} ( '\<equiv>' | '==' ) @{syntax term} ( '(' @'unchecked' ')' )? |
69597 | 1049 |
\<close> |
27040 | 1050 |
|
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1051 |
\<^descr> \<^theory_text>\<open>consts c :: \<sigma>\<close> declares constant \<open>c\<close> to have any instance of type scheme |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1052 |
\<open>\<sigma>\<close>. The optional mixfix annotations may attach concrete syntax to the |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1053 |
constants declared. |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1054 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1055 |
\<^descr> \<^theory_text>\<open>overloading x\<^sub>1 \<equiv> c\<^sub>1 :: \<tau>\<^sub>1 \<dots> x\<^sub>n \<equiv> c\<^sub>n :: \<tau>\<^sub>n begin \<dots> end\<close> defines |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1056 |
a theory target (cf.\ \secref{sec:target}) which allows to specify already |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1057 |
declared constants via definitions in the body. These are identified by an |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1058 |
explicitly given mapping from variable names \<open>x\<^sub>i\<close> to constants \<open>c\<^sub>i\<close> at |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1059 |
particular type instances. The definitions themselves are established using |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1060 |
common specification tools, using the names \<open>x\<^sub>i\<close> as reference to the |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1061 |
corresponding constants. |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1062 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1063 |
Option \<^theory_text>\<open>(unchecked)\<close> disables global dependency checks for the |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1064 |
corresponding definition, which is occasionally useful for exotic |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1065 |
overloading; this is a form of axiomatic specification. It is at the |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1066 |
discretion of the user to avoid malformed theory specifications! |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1067 |
\<close> |
27040 | 1068 |
|
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1069 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1070 |
subsubsection \<open>Example\<close> |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1071 |
|
62173 | 1072 |
consts Length :: "'a \<Rightarrow> nat" |
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1073 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1074 |
overloading |
62173 | 1075 |
Length\<^sub>0 \<equiv> "Length :: unit \<Rightarrow> nat" |
1076 |
Length\<^sub>1 \<equiv> "Length :: 'a \<times> unit \<Rightarrow> nat" |
|
1077 |
Length\<^sub>2 \<equiv> "Length :: 'a \<times> 'b \<times> unit \<Rightarrow> nat" |
|
1078 |
Length\<^sub>3 \<equiv> "Length :: 'a \<times> 'b \<times> 'c \<times> unit \<Rightarrow> nat" |
|
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1079 |
begin |
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1080 |
|
62173 | 1081 |
fun Length\<^sub>0 :: "unit \<Rightarrow> nat" where "Length\<^sub>0 () = 0" |
1082 |
fun Length\<^sub>1 :: "'a \<times> unit \<Rightarrow> nat" where "Length\<^sub>1 (a, ()) = 1" |
|
1083 |
fun Length\<^sub>2 :: "'a \<times> 'b \<times> unit \<Rightarrow> nat" where "Length\<^sub>2 (a, b, ()) = 2" |
|
1084 |
fun Length\<^sub>3 :: "'a \<times> 'b \<times> 'c \<times> unit \<Rightarrow> nat" where "Length\<^sub>3 (a, b, c, ()) = 3" |
|
62172
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1085 |
|
7eaeae127955
updated section on "Overloaded constant definitions";
wenzelm
parents:
61891
diff
changeset
|
1086 |
end |
27040 | 1087 |
|
62173 | 1088 |
lemma "Length (a, b, c, ()) = 3" by simp |
1089 |
lemma "Length ((a, b), (c, d), ()) = 2" by simp |
|
1090 |
lemma "Length ((a, b, c, d, e), ()) = 1" by simp |
|
1091 |
||
27040 | 1092 |
|
58618 | 1093 |
section \<open>Incorporating ML code \label{sec:ML}\<close> |
27040 | 1094 |
|
58618 | 1095 |
text \<open> |
27040 | 1096 |
\begin{matharray}{rcl} |
62862 | 1097 |
@{command_def "SML_file"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
62903 | 1098 |
@{command_def "SML_file_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
1099 |
@{command_def "SML_file_no_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
61493 | 1100 |
@{command_def "ML_file"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
62903 | 1101 |
@{command_def "ML_file_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
1102 |
@{command_def "ML_file_no_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
61493 | 1103 |
@{command_def "ML"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
68276 | 1104 |
@{command_def "ML_export"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
61493 | 1105 |
@{command_def "ML_prf"} & : & \<open>proof \<rightarrow> proof\<close> \\ |
1106 |
@{command_def "ML_val"} & : & \<open>any \<rightarrow>\<close> \\ |
|
1107 |
@{command_def "ML_command"} & : & \<open>any \<rightarrow>\<close> \\ |
|
1108 |
@{command_def "setup"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
1109 |
@{command_def "local_setup"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
1110 |
@{command_def "attribute_setup"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28758
diff
changeset
|
1111 |
\end{matharray} |
57478 | 1112 |
\begin{tabular}{rcll} |
61493 | 1113 |
@{attribute_def ML_print_depth} & : & \<open>attribute\<close> & default 10 \\ |
1114 |
@{attribute_def ML_source_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
|
62903 | 1115 |
@{attribute_def ML_debugger} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
61493 | 1116 |
@{attribute_def ML_exception_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
62903 | 1117 |
@{attribute_def ML_exception_debugger} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
70260 | 1118 |
@{attribute_def ML_environment} & : & \<open>attribute\<close> & default \<open>Isabelle\<close> \\ |
57478 | 1119 |
\end{tabular} |
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28758
diff
changeset
|
1120 |
|
69597 | 1121 |
\<^rail>\<open> |
62903 | 1122 |
(@@{command SML_file} | |
1123 |
@@{command SML_file_debug} | |
|
1124 |
@@{command SML_file_no_debug} | |
|
1125 |
@@{command ML_file} | |
|
1126 |
@@{command ML_file_debug} | |
|
62969 | 1127 |
@@{command ML_file_no_debug}) @{syntax name} ';'? |
27040 | 1128 |
; |
68276 | 1129 |
(@@{command ML} | @@{command ML_export} | @@{command ML_prf} | |
1130 |
@@{command ML_val} | @@{command ML_command} | @@{command setup} | |
|
1131 |
@@{command local_setup}) @{syntax text} |
|
27040 | 1132 |
; |
59783
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
wenzelm
parents:
59781
diff
changeset
|
1133 |
@@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}? |
69597 | 1134 |
\<close> |
27040 | 1135 |
|
61665 | 1136 |
\<^descr> \<^theory_text>\<open>SML_file name\<close> reads and evaluates the given Standard ML file. Top-level |
62862 | 1137 |
SML bindings are stored within the (global or local) theory context; the |
1138 |
initial environment is restricted to the Standard ML implementation of |
|
1139 |
Poly/ML, without the many add-ons of Isabelle/ML. Multiple \<^theory_text>\<open>SML_file\<close> |
|
1140 |
commands may be used to build larger Standard ML projects, independently of |
|
1141 |
the regular Isabelle/ML environment. |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
55837
diff
changeset
|
1142 |
|
61665 | 1143 |
\<^descr> \<^theory_text>\<open>ML_file name\<close> reads and evaluates the given ML file. The current theory |
69597 | 1144 |
context is passed down to the ML toplevel and may be modified, using \<^ML>\<open>Context.>>\<close> or derived ML commands. Top-level ML bindings are stored |
61665 | 1145 |
within the (global or local) theory context. |
62903 | 1146 |
|
1147 |
\<^descr> \<^theory_text>\<open>SML_file_debug\<close>, \<^theory_text>\<open>SML_file_no_debug\<close>, \<^theory_text>\<open>ML_file_debug\<close>, and |
|
1148 |
\<^theory_text>\<open>ML_file_no_debug\<close> change the @{attribute ML_debugger} option locally while |
|
1149 |
the given file is compiled. |
|
1150 |
||
68276 | 1151 |
\<^descr> \<^theory_text>\<open>ML\<close> is similar to \<^theory_text>\<open>ML_file\<close>, but evaluates directly the given \<open>text\<close>. |
1152 |
Top-level ML bindings are stored within the (global or local) theory |
|
61665 | 1153 |
context. |
28281 | 1154 |
|
68276 | 1155 |
\<^descr> \<^theory_text>\<open>ML_export\<close> is similar to \<^theory_text>\<open>ML\<close>, but the resulting toplevel bindings are |
1156 |
exported to the global bootstrap environment of the ML process --- it has |
|
68278 | 1157 |
a lasting effect that cannot be retracted. This allows ML evaluation |
68276 | 1158 |
without a formal theory context, e.g. for command-line tools via @{tool |
76987 | 1159 |
process} \<^cite>\<open>"isabelle-system"\<close>. |
68276 | 1160 |
|
61665 | 1161 |
\<^descr> \<^theory_text>\<open>ML_prf\<close> is analogous to \<^theory_text>\<open>ML\<close> but works within a proof context. |
1162 |
Top-level ML bindings are stored within the proof context in a purely |
|
1163 |
sequential fashion, disregarding the nested proof structure. ML bindings |
|
1164 |
introduced by \<^theory_text>\<open>ML_prf\<close> are discarded at the end of the proof. |
|
27040 | 1165 |
|
61665 | 1166 |
\<^descr> \<^theory_text>\<open>ML_val\<close> and \<^theory_text>\<open>ML_command\<close> are diagnostic versions of \<^theory_text>\<open>ML\<close>, which means |
1167 |
that the context may not be updated. \<^theory_text>\<open>ML_val\<close> echos the bindings produced |
|
1168 |
at the ML toplevel, but \<^theory_text>\<open>ML_command\<close> is silent. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
1169 |
|
61665 | 1170 |
\<^descr> \<^theory_text>\<open>setup "text"\<close> changes the current theory context by applying \<open>text\<close>, |
69597 | 1171 |
which refers to an ML expression of type \<^ML_type>\<open>theory -> theory\<close>. This |
61665 | 1172 |
enables to initialize any object-logic specific tools and packages written |
1173 |
in ML, for example. |
|
30461 | 1174 |
|
61665 | 1175 |
\<^descr> \<^theory_text>\<open>local_setup\<close> is similar to \<^theory_text>\<open>setup\<close> for a local theory context, and an |
69597 | 1176 |
ML expression of type \<^ML_type>\<open>local_theory -> local_theory\<close>. This allows |
61665 | 1177 |
to invoke local theory specification packages without going through concrete |
1178 |
outer syntax, for example. |
|
28758 | 1179 |
|
61665 | 1180 |
\<^descr> \<^theory_text>\<open>attribute_setup name = "text" description\<close> defines an attribute in the |
1181 |
current context. The given \<open>text\<close> has to be an ML expression of type |
|
69597 | 1182 |
\<^ML_type>\<open>attribute context_parser\<close>, cf.\ basic parsers defined in |
1183 |
structure \<^ML_structure>\<open>Args\<close> and \<^ML_structure>\<open>Attrib\<close>. |
|
30526 | 1184 |
|
1185 |
In principle, attributes can operate both on a given theorem and the |
|
61664 | 1186 |
implicit context, although in practice only one is modified and the other |
1187 |
serves as parameter. Here are examples for these two cases: |
|
58618 | 1188 |
\<close> |
30526 | 1189 |
|
59905 | 1190 |
(*<*)experiment begin(*>*) |
62903 | 1191 |
attribute_setup my_rule = |
1192 |
\<open>Attrib.thms >> (fn ths => |
|
1193 |
Thm.rule_attribute ths |
|
1194 |
(fn context: Context.generic => fn th: thm => |
|
1195 |
let val th' = th OF ths |
|
1196 |
in th' end))\<close> |
|
30526 | 1197 |
|
62903 | 1198 |
attribute_setup my_declaration = |
1199 |
\<open>Attrib.thms >> (fn ths => |
|
1200 |
Thm.declaration_attribute |
|
1201 |
(fn th: thm => fn context: Context.generic => |
|
1202 |
let val context' = context |
|
1203 |
in context' end))\<close> |
|
59905 | 1204 |
(*<*)end(*>*) |
30526 | 1205 |
|
58618 | 1206 |
text \<open> |
61664 | 1207 |
\<^descr> @{attribute ML_print_depth} controls the printing depth of the ML toplevel |
62903 | 1208 |
pretty printer. Typically the limit should be less than 10. Bigger values |
1209 |
such as 100--1000 are occasionally useful for debugging. |
|
57478 | 1210 |
|
61664 | 1211 |
\<^descr> @{attribute ML_source_trace} indicates whether the source text that is |
1212 |
given to the ML compiler should be output: it shows the raw Standard ML |
|
57478 | 1213 |
after expansion of Isabelle/ML antiquotations. |
1214 |
||
62903 | 1215 |
\<^descr> @{attribute ML_debugger} controls compilation of sources with or without |
1216 |
debugging information. The global system option @{system_option_ref |
|
1217 |
ML_debugger} does the same when building a session image. It is also |
|
1218 |
possible use commands like \<^theory_text>\<open>ML_file_debug\<close> etc. The ML debugger is |
|
76987 | 1219 |
explained further in \<^cite>\<open>"isabelle-jedit"\<close>. |
62903 | 1220 |
|
61664 | 1221 |
\<^descr> @{attribute ML_exception_trace} indicates whether the ML run-time system |
1222 |
should print a detailed stack trace on exceptions. The result is dependent |
|
62903 | 1223 |
on various ML compiler optimizations. The boundary for the exception trace |
1224 |
is the current Isar command transactions: it is occasionally better to |
|
69597 | 1225 |
insert the combinator \<^ML>\<open>Runtime.exn_trace\<close> into ML code for debugging |
76987 | 1226 |
\<^cite>\<open>"isabelle-implementation"\<close>, closer to the point where it actually |
62903 | 1227 |
happens. |
57478 | 1228 |
|
62903 | 1229 |
\<^descr> @{attribute ML_exception_debugger} controls detailed exception trace via |
1230 |
the Poly/ML debugger, at the cost of extra compile-time and run-time |
|
1231 |
overhead. Relevant ML modules need to be compiled beforehand with debugging |
|
1232 |
enabled, see @{attribute ML_debugger} above. |
|
70260 | 1233 |
|
1234 |
\<^descr> @{attribute ML_environment} determines the named ML environment for |
|
1235 |
toplevel declarations, e.g.\ in command \<^theory_text>\<open>ML\<close> or \<^theory_text>\<open>ML_file\<close>. The following |
|
1236 |
ML environments are predefined in Isabelle/Pure: |
|
1237 |
||
1238 |
\<^item> \<open>Isabelle\<close> for Isabelle/ML. It contains all modules of Isabelle/Pure and |
|
1239 |
further add-ons, e.g. material from Isabelle/HOL. |
|
1240 |
||
1241 |
\<^item> \<open>SML\<close> for official Standard ML. It contains only the initial basis |
|
1242 |
according to \<^url>\<open>http://sml-family.org/Basis/overview.html\<close>. |
|
1243 |
||
1244 |
The Isabelle/ML function \<^ML>\<open>ML_Env.setup\<close> defines a new ML environment. |
|
1245 |
This is useful to incorporate big SML projects in an isolated name space, |
|
1246 |
possibly with variations on ML syntax; the existing setup of |
|
1247 |
\<^ML>\<open>ML_Env.SML_operations\<close> follows the official standard. |
|
1248 |
||
1249 |
It is also possible to move toplevel bindings between ML environments, using |
|
1250 |
a notation with ``\<open>>\<close>'' as separator. For example: |
|
58618 | 1251 |
\<close> |
57478 | 1252 |
|
70260 | 1253 |
(*<*)experiment begin(*>*) |
1254 |
declare [[ML_environment = "Isabelle>SML"]] |
|
1255 |
ML \<open>val println = writeln\<close> |
|
1256 |
||
1257 |
declare [[ML_environment = "SML"]] |
|
1258 |
ML \<open>println "test"\<close> |
|
1259 |
||
1260 |
declare [[ML_environment = "Isabelle"]] |
|
1261 |
ML \<open>ML \<open>println\<close> (*bad*) handle ERROR msg => warning msg\<close> |
|
1262 |
(*<*)end(*>*) |
|
1263 |
||
27040 | 1264 |
|
70261 | 1265 |
section \<open>Generated files and exported files\<close> |
66757 | 1266 |
|
1267 |
text \<open> |
|
70057 | 1268 |
Write access to the physical file-system is incompatible with the stateless |
70261 | 1269 |
model of processing Isabelle documents. To avoid bad effects, the following |
1270 |
concepts for abstract file-management are provided by Isabelle: |
|
1271 |
||
1272 |
\<^descr>[Generated files] are stored within the theory context in Isabelle/ML. |
|
1273 |
This allows to operate on the content in Isabelle/ML, e.g. via the command |
|
1274 |
@{command compile_generated_files}. |
|
1275 |
||
1276 |
\<^descr>[Exported files] are stored within the session database in |
|
1277 |
Isabelle/Scala. This allows to deliver artefacts to external tools, see |
|
76987 | 1278 |
also \<^cite>\<open>"isabelle-system"\<close> for session \<^verbatim>\<open>ROOT\<close> declaration |
70261 | 1279 |
\<^theory_text>\<open>export_files\<close>, and @{tool build} option \<^verbatim>\<open>-e\<close>. |
1280 |
||
1281 |
A notable example is the command @{command_ref export_code} |
|
1282 |
(\chref{ch:export-code}): it uses both concepts simultaneously. |
|
1283 |
||
1284 |
File names are hierarchically structured, using a slash as separator. The |
|
1285 |
(long) theory name is used as a prefix: the resulting name needs to be |
|
1286 |
globally unique. |
|
70057 | 1287 |
|
1288 |
\begin{matharray}{rcll} |
|
1289 |
@{command_def "generate_file"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
1290 |
@{command_def "export_generated_files"} & : & \<open>context \<rightarrow>\<close> \\ |
|
1291 |
@{command_def "compile_generated_files"} & : & \<open>context \<rightarrow>\<close> \\ |
|
66757 | 1292 |
@{command_def "external_file"} & : & \<open>any \<rightarrow> any\<close> \\ |
1293 |
\end{matharray} |
|
1294 |
||
70057 | 1295 |
\<^rail>\<open> |
1296 |
@@{command generate_file} path '=' content |
|
1297 |
; |
|
1298 |
path: @{syntax embedded} |
|
1299 |
; |
|
1300 |
content: @{syntax embedded} |
|
1301 |
; |
|
1302 |
@@{command export_generated_files} (files_in_theory + @'and') |
|
1303 |
; |
|
1304 |
files_in_theory: (@'_' | (path+)) (('(' @'in' @{syntax name} ')')?) |
|
1305 |
; |
|
1306 |
@@{command compile_generated_files} (files_in_theory + @'and') \<newline> |
|
1307 |
(@'external_files' (external_files + @'and'))? \<newline> |
|
1308 |
(@'export_files' (export_files + @'and'))? \<newline> |
|
1309 |
(@'export_prefix' path)? |
|
1310 |
; |
|
1311 |
external_files: (path+) (('(' @'in' path ')')?) |
|
1312 |
; |
|
1313 |
export_files: (path+) (executable?) |
|
1314 |
; |
|
1315 |
executable: '(' ('exe' | 'executable') ')' |
|
1316 |
; |
|
1317 |
@@{command external_file} @{syntax name} ';'? |
|
1318 |
\<close> |
|
1319 |
||
1320 |
\<^descr> \<^theory_text>\<open>generate_file path = content\<close> augments the table of generated files |
|
1321 |
within the current theory by a new entry: duplicates are not allowed. The |
|
1322 |
name extension determines a pre-existent file-type; the \<open>content\<close> is a |
|
1323 |
string that is preprocessed according to rules of this file-type. |
|
1324 |
||
1325 |
For example, Isabelle/Pure supports \<^verbatim>\<open>.hs\<close> as file-type for Haskell: |
|
1326 |
embedded cartouches are evaluated as Isabelle/ML expressions of type |
|
1327 |
\<^ML_type>\<open>string\<close>, the result is inlined in Haskell string syntax. |
|
1328 |
||
1329 |
\<^descr> \<^theory_text>\<open>export_generated_files paths (in thy)\<close> retrieves named generated files |
|
1330 |
from the given theory (that needs be reachable via imports of the current |
|
1331 |
one). By default, the current theory node is used. Using ``\<^verbatim>\<open>_\<close>'' |
|
1332 |
(underscore) instead of explicit path names refers to \emph{all} files of a |
|
1333 |
theory node. |
|
1334 |
||
1335 |
The overall list of files is prefixed with the respective (long) theory name |
|
1336 |
and exported to the session database. In Isabelle/jEdit the result can be |
|
1337 |
browsed via the virtual file-system with prefix ``\<^verbatim>\<open>isabelle-export:\<close>'' |
|
1338 |
(using the regular file-browser). |
|
1339 |
||
75689 | 1340 |
\<^descr> \<^theory_text>\<open>scala_build_generated_files paths (in thy)\<close> retrieves named generated |
1341 |
files as for \<^theory_text>\<open>export_generated_files\<close> and writes them into a temporary |
|
1342 |
directory, which is taken as starting point for build process of |
|
76987 | 1343 |
Isabelle/Scala/Java modules (see \<^cite>\<open>"isabelle-system"\<close>). The |
76923 | 1344 |
corresponding @{path \<open>build.props\<close>} file is expected directly in the toplevel |
1345 |
directory, instead of @{path \<open>etc/build.props\<close>} for Isabelle system |
|
75689 | 1346 |
components. These properties need to specify sources, resources, services |
76110
0605eb327e60
more documentation of 'export_classpath' (session ROOT) and 'scala_build_generated_files' (Isar command);
wenzelm
parents:
75689
diff
changeset
|
1347 |
etc. as usual. The resulting JAR module becomes an export artefact of the |
0605eb327e60
more documentation of 'export_classpath' (session ROOT) and 'scala_build_generated_files' (Isar command);
wenzelm
parents:
75689
diff
changeset
|
1348 |
session database, with a name of the form |
0605eb327e60
more documentation of 'export_classpath' (session ROOT) and 'scala_build_generated_files' (Isar command);
wenzelm
parents:
75689
diff
changeset
|
1349 |
``\<open>theory\<close>\<^verbatim>\<open>:classpath/\<close>\<open>module\<close>\<^verbatim>\<open>.jar\<close>''. |
75689 | 1350 |
|
70057 | 1351 |
\<^descr> \<^theory_text>\<open>compile_generated_files paths (in thy) where compile_body\<close> retrieves |
1352 |
named generated files as for \<^theory_text>\<open>export_generated_files\<close> and writes them into |
|
1353 |
a temporary directory, such that the \<open>compile_body\<close> may operate on them as |
|
1354 |
an ML function of type \<^ML_type>\<open>Path.T -> unit\<close>. This may create further |
|
1355 |
files, e.g.\ executables produced by a compiler that is invoked as external |
|
1356 |
process (e.g.\ via \<^ML>\<open>Isabelle_System.bash\<close>), or any other files. |
|
1357 |
||
1358 |
The option ``\<^theory_text>\<open>external_files paths (in base_dir)\<close>'' copies files from the |
|
1359 |
physical file-system into the temporary directory, \emph{before} invoking |
|
1360 |
\<open>compile_body\<close>. The \<open>base_dir\<close> prefix is removed from each of the \<open>paths\<close>, |
|
1361 |
but the remaining sub-directory structure is reconstructed in the target |
|
1362 |
directory. |
|
1363 |
||
1364 |
The option ``\<^theory_text>\<open>export_files paths\<close>'' exports the specified files from the |
|
1365 |
temporary directory to the session database, \emph{after} invoking |
|
1366 |
\<open>compile_body\<close>. Entries may be decorated with ``\<^theory_text>\<open>(exe)\<close>'' to say that it is |
|
1367 |
a platform-specific executable program: the executable file-attribute will |
|
1368 |
be set, and on Windows the \<^verbatim>\<open>.exe\<close> file-extension will be included; |
|
1369 |
``\<^theory_text>\<open>(executable)\<close>'' only refers to the file-attribute, without special |
|
1370 |
treatment of the \<^verbatim>\<open>.exe\<close> extension. |
|
1371 |
||
1372 |
The option ``\<^theory_text>\<open>export_prefix path\<close>'' specifies an extra path prefix for all |
|
1373 |
exports of \<^theory_text>\<open>export_files\<close> above. |
|
66757 | 1374 |
|
1375 |
\<^descr> \<^theory_text>\<open>external_file name\<close> declares the formal dependency on the given file |
|
76987 | 1376 |
name, such that the Isabelle build process knows about it (see also \<^cite>\<open>"isabelle-system"\<close>). This is required for any files mentioned in |
70057 | 1377 |
\<^theory_text>\<open>compile_generated_files / external_files\<close> above, in order to document |
1378 |
source dependencies properly. It is also possible to use \<^theory_text>\<open>external_file\<close> |
|
1379 |
alone, e.g.\ when other Isabelle/ML tools use \<^ML>\<open>File.read\<close>, without |
|
1380 |
specific management of content by the Prover IDE. |
|
66757 | 1381 |
\<close> |
1382 |
||
1383 |
||
58618 | 1384 |
section \<open>Primitive specification elements\<close> |
27040 | 1385 |
|
58618 | 1386 |
subsection \<open>Sorts\<close> |
27040 | 1387 |
|
58618 | 1388 |
text \<open> |
27040 | 1389 |
\begin{matharray}{rcll} |
61493 | 1390 |
@{command_def "default_sort"} & : & \<open>local_theory \<rightarrow> local_theory\<close> |
27040 | 1391 |
\end{matharray} |
1392 |
||
69597 | 1393 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
1394 |
@@{command default_sort} @{syntax sort} |
69597 | 1395 |
\<close> |
27040 | 1396 |
|
61665 | 1397 |
\<^descr> \<^theory_text>\<open>default_sort s\<close> makes sort \<open>s\<close> the new default sort for any type |
1398 |
variable that is given explicitly in the text, but lacks a sort constraint |
|
1399 |
(wrt.\ the current context). Type variables generated by type inference are |
|
1400 |
not affected. |
|
28767 | 1401 |
|
61664 | 1402 |
Usually the default sort is only changed when defining a new object-logic. |
69597 | 1403 |
For example, the default sort in Isabelle/HOL is \<^class>\<open>type\<close>, the class of |
61664 | 1404 |
all HOL types. |
28767 | 1405 |
|
61664 | 1406 |
When merging theories, the default sorts of the parents are logically |
1407 |
intersected, i.e.\ the representations as lists of classes are joined. |
|
58618 | 1408 |
\<close> |
27040 | 1409 |
|
1410 |
||
58618 | 1411 |
subsection \<open>Types \label{sec:types-pure}\<close> |
27040 | 1412 |
|
58618 | 1413 |
text \<open> |
27040 | 1414 |
\begin{matharray}{rcll} |
61493 | 1415 |
@{command_def "type_synonym"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
1416 |
@{command_def "typedecl"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
27040 | 1417 |
\end{matharray} |
1418 |
||
69597 | 1419 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
1420 |
@@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?) |
27040 | 1421 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
1422 |
@@{command typedecl} @{syntax typespec} @{syntax mixfix}? |
69597 | 1423 |
\<close> |
27040 | 1424 |
|
61665 | 1425 |
\<^descr> \<^theory_text>\<open>type_synonym (\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = \<tau>\<close> introduces a \<^emph>\<open>type synonym\<close> \<open>(\<alpha>\<^sub>1, \<dots>, |
1426 |
\<alpha>\<^sub>n) t\<close> for the existing type \<open>\<tau>\<close>. Unlike the semantic type definitions in |
|
1427 |
Isabelle/HOL, type synonyms are merely syntactic abbreviations without any |
|
1428 |
logical significance. Internally, type synonyms are fully expanded. |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
63680
diff
changeset
|
1429 |
|
61665 | 1430 |
\<^descr> \<^theory_text>\<open>typedecl (\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t\<close> declares a new type constructor \<open>t\<close>. If the |
1431 |
object-logic defines a base sort \<open>s\<close>, then the constructor is declared to |
|
1432 |
operate on that, via the axiomatic type-class instance \<open>t :: (s, \<dots>, s)s\<close>. |
|
27040 | 1433 |
|
57487 | 1434 |
|
1435 |
\begin{warn} |
|
61664 | 1436 |
If you introduce a new type axiomatically, i.e.\ via @{command_ref |
1437 |
typedecl} and @{command_ref axiomatization} |
|
1438 |
(\secref{sec:axiomatizations}), the minimum requirement is that it has a |
|
1439 |
non-empty model, to avoid immediate collapse of the logical environment. |
|
1440 |
Moreover, one needs to demonstrate that the interpretation of such |
|
1441 |
free-form axiomatizations can coexist with other axiomatization schemes |
|
1442 |
for types, notably @{command_def typedef} in Isabelle/HOL |
|
1443 |
(\secref{sec:hol-typedef}), or any other extension that people might have |
|
1444 |
introduced elsewhere. |
|
57487 | 1445 |
\end{warn} |
58618 | 1446 |
\<close> |
27040 | 1447 |
|
1448 |
||
58618 | 1449 |
section \<open>Naming existing theorems \label{sec:theorems}\<close> |
27040 | 1450 |
|
58618 | 1451 |
text \<open> |
27040 | 1452 |
\begin{matharray}{rcll} |
61493 | 1453 |
@{command_def "lemmas"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
1454 |
@{command_def "named_theorems"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
27040 | 1455 |
\end{matharray} |
1456 |
||
69597 | 1457 |
\<^rail>\<open> |
62969 | 1458 |
@@{command lemmas} (@{syntax thmdef}? @{syntax thms} + @'and') |
59785 | 1459 |
@{syntax for_fixes} |
57946
6a26aa5fa65e
updated documentation concerning 'named_theorems';
wenzelm
parents:
57941
diff
changeset
|
1460 |
; |
59785 | 1461 |
@@{command named_theorems} (@{syntax name} @{syntax text}? + @'and') |
69597 | 1462 |
\<close> |
27040 | 1463 |
|
61665 | 1464 |
\<^descr> \<^theory_text>\<open>lemmas a = b\<^sub>1 \<dots> b\<^sub>n\<close>~@{keyword_def "for"}~\<open>x\<^sub>1 \<dots> x\<^sub>m\<close> evaluates given |
1465 |
facts (with attributes) in the current context, which may be augmented by |
|
1466 |
local variables. Results are standardized before being stored, i.e.\ |
|
1467 |
schematic variables are renamed to enforce index \<open>0\<close> uniformly. |
|
45600
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents:
42936
diff
changeset
|
1468 |
|
61665 | 1469 |
\<^descr> \<^theory_text>\<open>named_theorems name description\<close> declares a dynamic fact within the |
1470 |
context. The same \<open>name\<close> is used to define an attribute with the usual |
|
1471 |
\<open>add\<close>/\<open>del\<close> syntax (e.g.\ see \secref{sec:simp-rules}) to maintain the |
|
61664 | 1472 |
content incrementally, in canonical declaration order of the text structure. |
58618 | 1473 |
\<close> |
27040 | 1474 |
|
1475 |
||
70560 | 1476 |
section \<open>Oracles \label{sec:oracles}\<close> |
27040 | 1477 |
|
58618 | 1478 |
text \<open> |
47483 | 1479 |
\begin{matharray}{rcll} |
61493 | 1480 |
@{command_def "oracle"} & : & \<open>theory \<rightarrow> theory\<close> & (axiomatic!) \\ |
70560 | 1481 |
@{command_def "thm_oracles"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
47483 | 1482 |
\end{matharray} |
1483 |
||
61664 | 1484 |
Oracles allow Isabelle to take advantage of external reasoners such as |
1485 |
arithmetic decision procedures, model checkers, fast tautology checkers or |
|
1486 |
computer algebra systems. Invoked as an oracle, an external reasoner can |
|
1487 |
create arbitrary Isabelle theorems. |
|
28756 | 1488 |
|
61664 | 1489 |
It is the responsibility of the user to ensure that the external reasoner is |
1490 |
as trustworthy as the application requires. Another typical source of errors |
|
1491 |
is the linkup between Isabelle and the external tool, not just its concrete |
|
1492 |
implementation, but also the required translation between two different |
|
1493 |
logical environments. |
|
28756 | 1494 |
|
1495 |
Isabelle merely guarantees well-formedness of the propositions being |
|
61664 | 1496 |
asserted, and records within the internal derivation object how presumed |
70560 | 1497 |
theorems depend on unproven suppositions. This also includes implicit |
1498 |
type-class reasoning via the order-sorted algebra of class relations and |
|
1499 |
type arities (see also @{command_ref instantiation} and @{command_ref |
|
1500 |
instance}). |
|
28756 | 1501 |
|
69597 | 1502 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
41435
diff
changeset
|
1503 |
@@{command oracle} @{syntax name} '=' @{syntax text} |
70560 | 1504 |
; |
1505 |
@@{command thm_oracles} @{syntax thms} |
|
69597 | 1506 |
\<close> |
27040 | 1507 |
|
61665 | 1508 |
\<^descr> \<^theory_text>\<open>oracle name = "text"\<close> turns the given ML expression \<open>text\<close> of type |
69597 | 1509 |
\<^ML_text>\<open>'a -> cterm\<close> into an ML function of type \<^ML_text>\<open>'a -> thm\<close>, |
1510 |
which is bound to the global identifier \<^ML_text>\<open>name\<close>. This acts like an |
|
61665 | 1511 |
infinitary specification of axioms! Invoking the oracle only works within |
1512 |
the scope of the resulting theory. |
|
27040 | 1513 |
|
71926 | 1514 |
See \<^file>\<open>~~/src/HOL/Examples/Iff_Oracle.thy\<close> for a worked example of defining |
1515 |
a new primitive rule as oracle, and turning it into a proof method. |
|
70560 | 1516 |
|
1517 |
\<^descr> \<^theory_text>\<open>thm_oracles thms\<close> displays all oracles used in the internal derivation |
|
1518 |
of the given theorems; this covers the full graph of transitive |
|
1519 |
dependencies. |
|
58618 | 1520 |
\<close> |
27040 | 1521 |
|
1522 |
||
58618 | 1523 |
section \<open>Name spaces\<close> |
27040 | 1524 |
|
58618 | 1525 |
text \<open> |
27040 | 1526 |
\begin{matharray}{rcl} |
66248 | 1527 |
@{command_def "alias"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
1528 |
@{command_def "type_alias"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
|
61493 | 1529 |
@{command_def "hide_class"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
1530 |
@{command_def "hide_type"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
1531 |
@{command_def "hide_const"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
1532 |
@{command_def "hide_fact"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
27040 | 1533 |
\end{matharray} |
1534 |
||
69597 | 1535 |
\<^rail>\<open> |
66248 | 1536 |
(@{command alias} | @{command type_alias}) @{syntax name} '=' @{syntax name} |
1537 |
; |
|
1538 |
(@{command hide_class} | @{command hide_type} | |
|
1539 |
@{command hide_const} | @{command hide_fact}) ('(' @'open' ')')? (@{syntax name} + ) |
|
69597 | 1540 |
\<close> |
27040 | 1541 |
|
61664 | 1542 |
Isabelle organizes any kind of name declarations (of types, constants, |
1543 |
theorems etc.) by separate hierarchically structured name spaces. Normally |
|
61665 | 1544 |
the user does not have to control the behaviour of name spaces by hand, yet |
61664 | 1545 |
the following commands provide some way to do so. |
27040 | 1546 |
|
66248 | 1547 |
\<^descr> \<^theory_text>\<open>alias\<close> and \<^theory_text>\<open>type_alias\<close> introduce aliases for constants and type |
1548 |
constructors, respectively. This allows adhoc changes to name-space |
|
1549 |
accesses. |
|
1550 |
||
1551 |
\<^descr> \<^theory_text>\<open>type_alias b = c\<close> introduces an alias for an existing type constructor. |
|
1552 |
||
61665 | 1553 |
\<^descr> \<^theory_text>\<open>hide_class names\<close> fully removes class declarations from a given name |
1554 |
space; with the \<open>(open)\<close> option, only the unqualified base name is hidden. |
|
36177
8e0770d2e499
separate commands 'hide_class', 'hide_type', 'hide_const', 'hide_fact';
wenzelm
parents:
35681
diff
changeset
|
1555 |
|
61664 | 1556 |
Note that hiding name space accesses has no impact on logical declarations |
1557 |
--- they remain valid internally. Entities that are no longer accessible to |
|
1558 |
the user are printed with the special qualifier ``\<open>??\<close>'' prefixed to the |
|
1559 |
full internal name. |
|
27040 | 1560 |
|
61665 | 1561 |
\<^descr> \<^theory_text>\<open>hide_type\<close>, \<^theory_text>\<open>hide_const\<close>, and \<^theory_text>\<open>hide_fact\<close> are similar to |
1562 |
\<^theory_text>\<open>hide_class\<close>, but hide types, constants, and facts, respectively. |
|
58618 | 1563 |
\<close> |
27040 | 1564 |
|
26869 | 1565 |
end |