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