| author | krauss | 
| Mon, 23 May 2011 10:58:21 +0200 | |
| changeset 42949 | 618adb3584e5 | 
| parent 42616 | 92715b528e78 | 
| child 43326 | 47cf4bc789aa | 
| permissions | -rw-r--r-- | 
| 29755 | 1 | theory Prelim | 
| 2 | imports Base | |
| 3 | begin | |
| 18537 | 4 | |
| 5 | chapter {* Preliminaries *}
 | |
| 6 | ||
| 20429 | 7 | section {* Contexts \label{sec:context} *}
 | 
| 18537 | 8 | |
| 20429 | 9 | text {*
 | 
| 20451 | 10 | A logical context represents the background that is required for | 
| 11 | formulating statements and composing proofs. It acts as a medium to | |
| 12 | produce formal content, depending on earlier material (declarations, | |
| 13 | results etc.). | |
| 18537 | 14 | |
| 20451 | 15 | For example, derivations within the Isabelle/Pure logic can be | 
| 16 |   described as a judgment @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"}, which means that a
 | |
| 20429 | 17 |   proposition @{text "\<phi>"} is derivable from hypotheses @{text "\<Gamma>"}
 | 
| 18 |   within the theory @{text "\<Theta>"}.  There are logical reasons for
 | |
| 20451 | 19 |   keeping @{text "\<Theta>"} and @{text "\<Gamma>"} separate: theories can be
 | 
| 20 | liberal about supporting type constructors and schematic | |
| 21 | polymorphism of constants and axioms, while the inner calculus of | |
| 22 |   @{text "\<Gamma> \<turnstile> \<phi>"} is strictly limited to Simple Type Theory (with
 | |
| 23 | fixed type variables in the assumptions). | |
| 18537 | 24 | |
| 20429 | 25 | \medskip Contexts and derivations are linked by the following key | 
| 26 | principles: | |
| 27 | ||
| 28 |   \begin{itemize}
 | |
| 29 | ||
| 30 | \item Transfer: monotonicity of derivations admits results to be | |
| 20451 | 31 |   transferred into a \emph{larger} context, i.e.\ @{text "\<Gamma> \<turnstile>\<^sub>\<Theta>
 | 
| 32 |   \<phi>"} implies @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>"} for contexts @{text "\<Theta>'
 | |
| 33 |   \<supseteq> \<Theta>"} and @{text "\<Gamma>' \<supseteq> \<Gamma>"}.
 | |
| 18537 | 34 | |
| 20429 | 35 | \item Export: discharge of hypotheses admits results to be exported | 
| 20451 | 36 |   into a \emph{smaller} context, i.e.\ @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>"}
 | 
| 37 |   implies @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>"} where @{text "\<Gamma>' \<supseteq> \<Gamma>"} and
 | |
| 38 |   @{text "\<Delta> = \<Gamma>' - \<Gamma>"}.  Note that @{text "\<Theta>"} remains unchanged here,
 | |
| 39 |   only the @{text "\<Gamma>"} part is affected.
 | |
| 20429 | 40 | |
| 41 |   \end{itemize}
 | |
| 18537 | 42 | |
| 20451 | 43 | \medskip By modeling the main characteristics of the primitive | 
| 44 |   @{text "\<Theta>"} and @{text "\<Gamma>"} above, and abstracting over any
 | |
| 45 | particular logical content, we arrive at the fundamental notions of | |
| 46 |   \emph{theory context} and \emph{proof context} in Isabelle/Isar.
 | |
| 47 |   These implement a certain policy to manage arbitrary \emph{context
 | |
| 48 | data}. There is a strongly-typed mechanism to declare new kinds of | |
| 20429 | 49 | data at compile time. | 
| 18537 | 50 | |
| 20451 | 51 | The internal bootstrap process of Isabelle/Pure eventually reaches a | 
| 52 |   stage where certain data slots provide the logical content of @{text
 | |
| 53 |   "\<Theta>"} and @{text "\<Gamma>"} sketched above, but this does not stop there!
 | |
| 54 | Various additional data slots support all kinds of mechanisms that | |
| 55 | are not necessarily part of the core logic. | |
| 18537 | 56 | |
| 20429 | 57 | For example, there would be data for canonical introduction and | 
| 58 | elimination rules for arbitrary operators (depending on the | |
| 59 | object-logic and application), which enables users to perform | |
| 20451 | 60 |   standard proof steps implicitly (cf.\ the @{text "rule"} method
 | 
| 61 |   \cite{isabelle-isar-ref}).
 | |
| 18537 | 62 | |
| 20451 | 63 | \medskip Thus Isabelle/Isar is able to bring forth more and more | 
| 64 | concepts successively. In particular, an object-logic like | |
| 65 | Isabelle/HOL continues the Isabelle/Pure setup by adding specific | |
| 66 | components for automated reasoning (classical reasoner, tableau | |
| 67 | prover, structured induction etc.) and derived specification | |
| 68 | mechanisms (inductive predicates, recursive functions etc.). All of | |
| 69 | this is ultimately based on the generic data management by theory | |
| 70 | and proof contexts introduced here. | |
| 18537 | 71 | *} | 
| 72 | ||
| 73 | ||
| 74 | subsection {* Theory context \label{sec:context-theory} *}
 | |
| 75 | ||
| 34921 | 76 | text {* A \emph{theory} is a data container with explicit name and
 | 
| 77 | unique identifier. Theories are related by a (nominal) sub-theory | |
| 20451 | 78 | relation, which corresponds to the dependency graph of the original | 
| 79 | construction; each theory is derived from a certain sub-graph of | |
| 34921 | 80 | ancestor theories. To this end, the system maintains a set of | 
| 81 | symbolic ``identification stamps'' within each theory. | |
| 18537 | 82 | |
| 34921 | 83 | In order to avoid the full-scale overhead of explicit sub-theory | 
| 84 | identification of arbitrary intermediate stages, a theory is | |
| 85 |   switched into @{text "draft"} mode under certain circumstances.  A
 | |
| 86 | draft theory acts like a linear type, where updates invalidate | |
| 87 |   earlier versions.  An invalidated draft is called \emph{stale}.
 | |
| 20429 | 88 | |
| 34921 | 89 |   The @{text "checkpoint"} operation produces a safe stepping stone
 | 
| 90 | that will survive the next update without becoming stale: both the | |
| 91 | old and the new theory remain valid and are related by the | |
| 92 | sub-theory relation. Checkpointing essentially recovers purely | |
| 93 | functional theory values, at the expense of some extra internal | |
| 94 | bookkeeping. | |
| 20447 | 95 | |
| 96 |   The @{text "copy"} operation produces an auxiliary version that has
 | |
| 97 | the same data content, but is unrelated to the original: updates of | |
| 98 | the copy do not affect the original, neither does the sub-theory | |
| 99 | relation hold. | |
| 20429 | 100 | |
| 34921 | 101 |   The @{text "merge"} operation produces the least upper bound of two
 | 
| 102 | theories, which actually degenerates into absorption of one theory | |
| 103 | into the other (according to the nominal sub-theory relation). | |
| 104 | ||
| 105 |   The @{text "begin"} operation starts a new theory by importing
 | |
| 106 | several parent theories and entering a special mode of nameless | |
| 107 |   incremental updates, until the final @{text "end"} operation is
 | |
| 108 | performed. | |
| 109 | ||
| 20447 | 110 |   \medskip The example in \figref{fig:ex-theory} below shows a theory
 | 
| 20451 | 111 |   graph derived from @{text "Pure"}, with theory @{text "Length"}
 | 
| 112 |   importing @{text "Nat"} and @{text "List"}.  The body of @{text
 | |
| 113 | "Length"} consists of a sequence of updates, working mostly on | |
| 34921 | 114 | drafts internally, while transaction boundaries of Isar top-level | 
| 115 |   commands (\secref{sec:isar-toplevel}) are guaranteed to be safe
 | |
| 116 | checkpoints. | |
| 20447 | 117 | |
| 118 |   \begin{figure}[htb]
 | |
| 119 |   \begin{center}
 | |
| 20429 | 120 |   \begin{tabular}{rcccl}
 | 
| 20447 | 121 |         &            & @{text "Pure"} \\
 | 
| 122 |         &            & @{text "\<down>"} \\
 | |
| 123 |         &            & @{text "FOL"} \\
 | |
| 18537 | 124 | & $\swarrow$ & & $\searrow$ & \\ | 
| 21852 | 125 |   @{text "Nat"} &    &              &            & @{text "List"} \\
 | 
| 18537 | 126 | & $\searrow$ & & $\swarrow$ \\ | 
| 20447 | 127 |         &            & @{text "Length"} \\
 | 
| 26864 | 128 |         &            & \multicolumn{3}{l}{~~@{keyword "imports"}} \\
 | 
| 129 |         &            & \multicolumn{3}{l}{~~@{keyword "begin"}} \\
 | |
| 18537 | 130 | & & $\vdots$~~ \\ | 
| 20447 | 131 |         &            & @{text "\<bullet>"}~~ \\
 | 
| 132 | & & $\vdots$~~ \\ | |
| 133 |         &            & @{text "\<bullet>"}~~ \\
 | |
| 134 | & & $\vdots$~~ \\ | |
| 26864 | 135 |         &            & \multicolumn{3}{l}{~~@{command "end"}} \\
 | 
| 20429 | 136 |   \end{tabular}
 | 
| 20451 | 137 |   \caption{A theory definition depending on ancestors}\label{fig:ex-theory}
 | 
| 20447 | 138 |   \end{center}
 | 
| 139 |   \end{figure}
 | |
| 20451 | 140 | |
| 141 |   \medskip There is a separate notion of \emph{theory reference} for
 | |
| 142 | maintaining a live link to an evolving theory context: updates on | |
| 39821 | 143 | drafts are propagated automatically. Dynamic updating stops when | 
| 144 |   the next @{text "checkpoint"} is reached.
 | |
| 20451 | 145 | |
| 146 | Derived entities may store a theory reference in order to indicate | |
| 39821 | 147 | the formal context from which they are derived. This implicitly | 
| 148 | assumes monotonic reasoning, because the referenced context may | |
| 149 | become larger without further notice. | |
| 18537 | 150 | *} | 
| 151 | ||
| 20430 | 152 | text %mlref {*
 | 
| 20447 | 153 |   \begin{mldecls}
 | 
| 154 |   @{index_ML_type theory} \\
 | |
| 39837 | 155 |   @{index_ML Theory.eq_thy: "theory * theory -> bool"} \\
 | 
| 20447 | 156 |   @{index_ML Theory.subthy: "theory * theory -> bool"} \\
 | 
| 157 |   @{index_ML Theory.checkpoint: "theory -> theory"} \\
 | |
| 20547 | 158 |   @{index_ML Theory.copy: "theory -> theory"} \\
 | 
| 34921 | 159 |   @{index_ML Theory.merge: "theory * theory -> theory"} \\
 | 
| 160 |   @{index_ML Theory.begin_theory: "string -> theory list -> theory"} \\
 | |
| 39837 | 161 |   @{index_ML Theory.parents_of: "theory -> theory list"} \\
 | 
| 162 |   @{index_ML Theory.ancestors_of: "theory -> theory list"} \\
 | |
| 20547 | 163 |   \end{mldecls}
 | 
| 164 |   \begin{mldecls}
 | |
| 20447 | 165 |   @{index_ML_type theory_ref} \\
 | 
| 166 |   @{index_ML Theory.deref: "theory_ref -> theory"} \\
 | |
| 24137 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 167 |   @{index_ML Theory.check_thy: "theory -> theory_ref"} \\
 | 
| 20447 | 168 |   \end{mldecls}
 | 
| 169 | ||
| 170 |   \begin{description}
 | |
| 171 | ||
| 39864 | 172 |   \item Type @{ML_type theory} represents theory contexts.  This is
 | 
| 39821 | 173 | essentially a linear type, with explicit runtime checking. | 
| 174 | Primitive theory operations destroy the original version, which then | |
| 175 | becomes ``stale''. This can be prevented by explicit checkpointing, | |
| 176 | which the system does at least at the boundary of toplevel command | |
| 177 |   transactions \secref{sec:isar-toplevel}.
 | |
| 20447 | 178 | |
| 39837 | 179 |   \item @{ML "Theory.eq_thy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} check strict
 | 
| 180 | identity of two theories. | |
| 181 | ||
| 34921 | 182 |   \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} compares theories
 | 
| 183 | according to the intrinsic graph structure of the construction. | |
| 184 | This sub-theory relation is a nominal approximation of inclusion | |
| 185 |   (@{text "\<subseteq>"}) of the corresponding content (according to the
 | |
| 186 | semantics of the ML modules that implement the data). | |
| 20447 | 187 | |
| 188 |   \item @{ML "Theory.checkpoint"}~@{text "thy"} produces a safe
 | |
| 34921 | 189 |   stepping stone in the linear development of @{text "thy"}.  This
 | 
| 190 | changes the old theory, but the next update will result in two | |
| 191 | related, valid theories. | |
| 20447 | 192 | |
| 193 |   \item @{ML "Theory.copy"}~@{text "thy"} produces a variant of @{text
 | |
| 34921 | 194 | "thy"} with the same data. The copy is not related to the original, | 
| 195 | but the original is unchanged. | |
| 196 | ||
| 197 |   \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} absorbs one theory
 | |
| 198 |   into the other, without changing @{text "thy\<^sub>1"} or @{text "thy\<^sub>2"}.
 | |
| 199 | This version of ad-hoc theory merge fails for unrelated theories! | |
| 200 | ||
| 201 |   \item @{ML "Theory.begin_theory"}~@{text "name parents"} constructs
 | |
| 39825 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 wenzelm parents: 
39821diff
changeset | 202 | a new theory based on the given parents. This ML function is | 
| 34921 | 203 | normally not invoked directly. | 
| 20447 | 204 | |
| 39837 | 205 |   \item @{ML "Theory.parents_of"}~@{text "thy"} returns the direct
 | 
| 206 |   ancestors of @{text thy}.
 | |
| 207 | ||
| 208 |   \item @{ML "Theory.ancestors_of"}~@{text "thy"} returns all
 | |
| 209 |   ancestors of @{text thy} (not including @{text thy} itself).
 | |
| 210 | ||
| 39864 | 211 |   \item Type @{ML_type theory_ref} represents a sliding reference to
 | 
| 212 | an always valid theory; updates on the original are propagated | |
| 20447 | 213 | automatically. | 
| 214 | ||
| 24137 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 215 |   \item @{ML "Theory.deref"}~@{text "thy_ref"} turns a @{ML_type
 | 
| 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 216 |   "theory_ref"} into an @{ML_type "theory"} value.  As the referenced
 | 
| 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 217 |   theory evolves monotonically over time, later invocations of @{ML
 | 
| 20451 | 218 | "Theory.deref"} may refer to a larger context. | 
| 20447 | 219 | |
| 24137 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 220 |   \item @{ML "Theory.check_thy"}~@{text "thy"} produces a @{ML_type
 | 
| 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 221 |   "theory_ref"} from a valid @{ML_type "theory"} value.
 | 
| 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
22869diff
changeset | 222 | |
| 20447 | 223 |   \end{description}
 | 
| 20430 | 224 | *} | 
| 225 | ||
| 39832 | 226 | text %mlantiq {*
 | 
| 227 |   \begin{matharray}{rcl}
 | |
| 228 |   @{ML_antiquotation_def "theory"} & : & @{text ML_antiquotation} \\
 | |
| 229 |   \end{matharray}
 | |
| 230 | ||
| 42510 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 231 |   @{rail "
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 232 |   @@{ML_antiquotation theory} nameref?
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 233 | "} | 
| 39832 | 234 | |
| 235 |   \begin{description}
 | |
| 236 | ||
| 237 |   \item @{text "@{theory}"} refers to the background theory of the
 | |
| 238 | current context --- as abstract value. | |
| 239 | ||
| 240 |   \item @{text "@{theory A}"} refers to an explicitly named ancestor
 | |
| 241 |   theory @{text "A"} of the background theory of the current context
 | |
| 242 | --- as abstract value. | |
| 243 | ||
| 244 |   \end{description}
 | |
| 245 | *} | |
| 246 | ||
| 18537 | 247 | |
| 248 | subsection {* Proof context \label{sec:context-proof} *}
 | |
| 249 | ||
| 34921 | 250 | text {* A proof context is a container for pure data with a
 | 
| 39821 | 251 |   back-reference to the theory from which it is derived.  The @{text
 | 
| 252 | "init"} operation creates a proof context from a given theory. | |
| 34921 | 253 | Modifications to draft theories are propagated to the proof context | 
| 254 |   as usual, but there is also an explicit @{text "transfer"} operation
 | |
| 255 | to force resynchronization with more substantial updates to the | |
| 256 | underlying theory. | |
| 20429 | 257 | |
| 34921 | 258 | Entities derived in a proof context need to record logical | 
| 20447 | 259 | requirements explicitly, since there is no separate context | 
| 34921 | 260 | identification or symbolic inclusion as for theories. For example, | 
| 261 |   hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
 | |
| 262 |   are recorded separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, just to
 | |
| 263 | make double sure. Results could still leak into an alien proof | |
| 264 | context due to programming errors, but Isabelle/Isar includes some | |
| 265 | extra validity checks in critical positions, notably at the end of a | |
| 266 | sub-proof. | |
| 20429 | 267 | |
| 20451 | 268 | Proof contexts may be manipulated arbitrarily, although the common | 
| 269 | discipline is to follow block structure as a mental model: a given | |
| 270 | context is extended consecutively, and results are exported back | |
| 34921 | 271 | into the original context. Note that an Isar proof state models | 
| 20451 | 272 | block-structured reasoning explicitly, using a stack of proof | 
| 34921 | 273 | contexts internally. For various technical reasons, the background | 
| 274 | theory of an Isar proof state must not be changed while the proof is | |
| 275 | still under construction! | |
| 18537 | 276 | *} | 
| 277 | ||
| 20449 | 278 | text %mlref {*
 | 
| 279 |   \begin{mldecls}
 | |
| 280 |   @{index_ML_type Proof.context} \\
 | |
| 42361 | 281 |   @{index_ML Proof_Context.init_global: "theory -> Proof.context"} \\
 | 
| 282 |   @{index_ML Proof_Context.theory_of: "Proof.context -> theory"} \\
 | |
| 283 |   @{index_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\
 | |
| 20449 | 284 |   \end{mldecls}
 | 
| 285 | ||
| 286 |   \begin{description}
 | |
| 287 | ||
| 39864 | 288 |   \item Type @{ML_type Proof.context} represents proof contexts.
 | 
| 289 | Elements of this type are essentially pure values, with a sliding | |
| 290 | reference to the background theory. | |
| 20449 | 291 | |
| 42361 | 292 |   \item @{ML Proof_Context.init_global}~@{text "thy"} produces a proof context
 | 
| 20449 | 293 |   derived from @{text "thy"}, initializing all data.
 | 
| 294 | ||
| 42361 | 295 |   \item @{ML Proof_Context.theory_of}~@{text "ctxt"} selects the
 | 
| 20451 | 296 |   background theory from @{text "ctxt"}, dereferencing its internal
 | 
| 297 |   @{ML_type theory_ref}.
 | |
| 20449 | 298 | |
| 42361 | 299 |   \item @{ML Proof_Context.transfer}~@{text "thy ctxt"} promotes the
 | 
| 20449 | 300 |   background theory of @{text "ctxt"} to the super theory @{text
 | 
| 301 | "thy"}. | |
| 302 | ||
| 303 |   \end{description}
 | |
| 304 | *} | |
| 305 | ||
| 39832 | 306 | text %mlantiq {*
 | 
| 307 |   \begin{matharray}{rcl}
 | |
| 308 |   @{ML_antiquotation_def "context"} & : & @{text ML_antiquotation} \\
 | |
| 309 |   \end{matharray}
 | |
| 310 | ||
| 311 |   \begin{description}
 | |
| 312 | ||
| 313 |   \item @{text "@{context}"} refers to \emph{the} context at
 | |
| 314 | compile-time --- as abstract value. Independently of (local) theory | |
| 315 | or proof mode, this always produces a meaningful result. | |
| 316 | ||
| 317 | This is probably the most common antiquotation in interactive | |
| 318 | experimentation with ML inside Isar. | |
| 319 | ||
| 320 |   \end{description}
 | |
| 321 | *} | |
| 322 | ||
| 20430 | 323 | |
| 20451 | 324 | subsection {* Generic contexts \label{sec:generic-context} *}
 | 
| 20429 | 325 | |
| 20449 | 326 | text {*
 | 
| 327 | A generic context is the disjoint sum of either a theory or proof | |
| 20451 | 328 | context. Occasionally, this enables uniform treatment of generic | 
| 20450 | 329 | context data, typically extra-logical information. Operations on | 
| 20449 | 330 | generic contexts include the usual injections, partial selections, | 
| 331 | and combinators for lifting operations on either component of the | |
| 332 | disjoint sum. | |
| 333 | ||
| 334 |   Moreover, there are total operations @{text "theory_of"} and @{text
 | |
| 335 | "proof_of"} to convert a generic context into either kind: a theory | |
| 20451 | 336 | can always be selected from the sum, while a proof context might | 
| 34921 | 337 |   have to be constructed by an ad-hoc @{text "init"} operation, which
 | 
| 338 | incurs a small runtime overhead. | |
| 20449 | 339 | *} | 
| 20430 | 340 | |
| 20449 | 341 | text %mlref {*
 | 
| 342 |   \begin{mldecls}
 | |
| 343 |   @{index_ML_type Context.generic} \\
 | |
| 344 |   @{index_ML Context.theory_of: "Context.generic -> theory"} \\
 | |
| 345 |   @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
 | |
| 346 |   \end{mldecls}
 | |
| 347 | ||
| 348 |   \begin{description}
 | |
| 20430 | 349 | |
| 39864 | 350 |   \item Type @{ML_type Context.generic} is the direct sum of @{ML_type
 | 
| 20451 | 351 |   "theory"} and @{ML_type "Proof.context"}, with the datatype
 | 
| 352 |   constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
 | |
| 20449 | 353 | |
| 354 |   \item @{ML Context.theory_of}~@{text "context"} always produces a
 | |
| 355 |   theory from the generic @{text "context"}, using @{ML
 | |
| 42361 | 356 | "Proof_Context.theory_of"} as required. | 
| 20449 | 357 | |
| 358 |   \item @{ML Context.proof_of}~@{text "context"} always produces a
 | |
| 359 |   proof context from the generic @{text "context"}, using @{ML
 | |
| 42361 | 360 | "Proof_Context.init_global"} as required (note that this re-initializes the | 
| 20451 | 361 | context data with each invocation). | 
| 20449 | 362 | |
| 363 |   \end{description}
 | |
| 364 | *} | |
| 20437 | 365 | |
| 20476 | 366 | |
| 367 | subsection {* Context data \label{sec:context-data} *}
 | |
| 20447 | 368 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 369 | text {* The main purpose of theory and proof contexts is to manage
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 370 | arbitrary (pure) data. New data types can be declared incrementally | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 371 | at compile time. There are separate declaration mechanisms for any | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 372 | of the three kinds of contexts: theory, proof, generic. | 
| 20449 | 373 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 374 |   \paragraph{Theory data} declarations need to implement the following
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 375 | SML signature: | 
| 20449 | 376 | |
| 377 | \medskip | |
| 378 |   \begin{tabular}{ll}
 | |
| 22869 | 379 |   @{text "\<type> T"} & representing type \\
 | 
| 380 |   @{text "\<val> empty: T"} & empty default value \\
 | |
| 381 |   @{text "\<val> extend: T \<rightarrow> T"} & re-initialize on import \\
 | |
| 382 |   @{text "\<val> merge: T \<times> T \<rightarrow> T"} & join on import \\
 | |
| 20449 | 383 |   \end{tabular}
 | 
| 384 | \medskip | |
| 385 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 386 |   The @{text "empty"} value acts as initial default for \emph{any}
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 387 |   theory that does not declare actual data content; @{text "extend"}
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 388 |   is acts like a unitary version of @{text "merge"}.
 | 
| 20449 | 389 | |
| 34921 | 390 |   Implementing @{text "merge"} can be tricky.  The general idea is
 | 
| 391 |   that @{text "merge (data\<^sub>1, data\<^sub>2)"} inserts those parts of @{text
 | |
| 392 |   "data\<^sub>2"} into @{text "data\<^sub>1"} that are not yet present, while
 | |
| 393 |   keeping the general order of things.  The @{ML Library.merge}
 | |
| 394 | function on plain lists may serve as canonical template. | |
| 395 | ||
| 396 | Particularly note that shared parts of the data must not be | |
| 397 | duplicated by naive concatenation, or a theory graph that is like a | |
| 398 | chain of diamonds would cause an exponential blowup! | |
| 399 | ||
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 400 |   \paragraph{Proof context data} declarations need to implement the
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 401 | following SML signature: | 
| 20449 | 402 | |
| 403 | \medskip | |
| 404 |   \begin{tabular}{ll}
 | |
| 22869 | 405 |   @{text "\<type> T"} & representing type \\
 | 
| 406 |   @{text "\<val> init: theory \<rightarrow> T"} & produce initial value \\
 | |
| 20449 | 407 |   \end{tabular}
 | 
| 408 | \medskip | |
| 409 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 410 |   The @{text "init"} operation is supposed to produce a pure value
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 411 | from the given background theory and should be somehow | 
| 34921 | 412 | ``immediate''. Whenever a proof context is initialized, which | 
| 413 |   happens frequently, the the system invokes the @{text "init"}
 | |
| 39821 | 414 |   operation of \emph{all} theory data slots ever declared.  This also
 | 
| 415 | means that one needs to be economic about the total number of proof | |
| 416 | data declarations in the system, i.e.\ each ML module should declare | |
| 417 | at most one, sometimes two data slots for its internal use. | |
| 418 | Repeated data declarations to simulate a record type should be | |
| 419 | avoided! | |
| 20449 | 420 | |
| 20451 | 421 |   \paragraph{Generic data} provides a hybrid interface for both theory
 | 
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 422 |   and proof data.  The @{text "init"} operation for proof contexts is
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 423 | predefined to select the current data value from the background | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 424 | theory. | 
| 20449 | 425 | |
| 39821 | 426 |   \bigskip Any of the above data declarations over type @{text "T"}
 | 
| 427 | result in an ML structure with the following signature: | |
| 20449 | 428 | |
| 429 | \medskip | |
| 430 |   \begin{tabular}{ll}
 | |
| 431 |   @{text "get: context \<rightarrow> T"} \\
 | |
| 432 |   @{text "put: T \<rightarrow> context \<rightarrow> context"} \\
 | |
| 433 |   @{text "map: (T \<rightarrow> T) \<rightarrow> context \<rightarrow> context"} \\
 | |
| 434 |   \end{tabular}
 | |
| 435 | \medskip | |
| 436 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 437 | These other operations provide exclusive access for the particular | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 438 | kind of context (theory, proof, or generic context). This interface | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 439 | observes the ML discipline for types and scopes: there is no other | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 440 | way to access the corresponding data slot of a context. By keeping | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 441 | these operations private, an Isabelle/ML module may maintain | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 442 | abstract values authentically. *} | 
| 20447 | 443 | |
| 20450 | 444 | text %mlref {*
 | 
| 445 |   \begin{mldecls}
 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 446 |   @{index_ML_functor Theory_Data} \\
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 447 |   @{index_ML_functor Proof_Data} \\
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 448 |   @{index_ML_functor Generic_Data} \\
 | 
| 20450 | 449 |   \end{mldecls}
 | 
| 450 | ||
| 451 |   \begin{description}
 | |
| 452 | ||
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 453 |   \item @{ML_functor Theory_Data}@{text "(spec)"} declares data for
 | 
| 20450 | 454 |   type @{ML_type theory} according to the specification provided as
 | 
| 20451 | 455 | argument structure. The resulting structure provides data init and | 
| 456 | access operations as described above. | |
| 20450 | 457 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 458 |   \item @{ML_functor Proof_Data}@{text "(spec)"} is analogous to
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 459 |   @{ML_functor Theory_Data} for type @{ML_type Proof.context}.
 | 
| 20450 | 460 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 461 |   \item @{ML_functor Generic_Data}@{text "(spec)"} is analogous to
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 462 |   @{ML_functor Theory_Data} for type @{ML_type Context.generic}.
 | 
| 20450 | 463 | |
| 464 |   \end{description}
 | |
| 465 | *} | |
| 466 | ||
| 34928 | 467 | text %mlex {*
 | 
| 468 | The following artificial example demonstrates theory | |
| 469 | data: we maintain a set of terms that are supposed to be wellformed | |
| 470 | wrt.\ the enclosing theory. The public interface is as follows: | |
| 471 | *} | |
| 472 | ||
| 473 | ML {*
 | |
| 474 | signature WELLFORMED_TERMS = | |
| 475 | sig | |
| 476 | val get: theory -> term list | |
| 477 | val add: term -> theory -> theory | |
| 478 | end; | |
| 479 | *} | |
| 480 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 481 | text {* The implementation uses private theory data internally, and
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 482 | only exposes an operation that involves explicit argument checking | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 483 | wrt.\ the given theory. *} | 
| 34928 | 484 | |
| 485 | ML {*
 | |
| 486 | structure Wellformed_Terms: WELLFORMED_TERMS = | |
| 487 | struct | |
| 488 | ||
| 489 | structure Terms = Theory_Data | |
| 490 | ( | |
| 39687 | 491 | type T = term Ord_List.T; | 
| 34928 | 492 | val empty = []; | 
| 493 | val extend = I; | |
| 494 | fun merge (ts1, ts2) = | |
| 39687 | 495 | Ord_List.union Term_Ord.fast_term_ord ts1 ts2; | 
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 496 | ); | 
| 34928 | 497 | |
| 498 | val get = Terms.get; | |
| 499 | ||
| 500 | fun add raw_t thy = | |
| 39821 | 501 | let | 
| 502 | val t = Sign.cert_term thy raw_t; | |
| 503 | in | |
| 504 | Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy | |
| 505 | end; | |
| 34928 | 506 | |
| 507 | end; | |
| 508 | *} | |
| 509 | ||
| 39864 | 510 | text {* Type @{ML_type "term Ord_List.T"} is used for reasonably
 | 
| 511 | efficient representation of a set of terms: all operations are | |
| 512 | linear in the number of stored elements. Here we assume that users | |
| 513 | of this module do not care about the declaration order, since that | |
| 514 | data structure forces its own arrangement of elements. | |
| 34928 | 515 | |
| 40153 | 516 |   Observe how the @{ML_text merge} operation joins the data slots of
 | 
| 39687 | 517 |   the two constituents: @{ML Ord_List.union} prevents duplication of
 | 
| 34928 | 518 | common data from different branches, thus avoiding the danger of | 
| 39821 | 519 | exponential blowup. Plain list append etc.\ must never be used for | 
| 520 | theory data merges! | |
| 34928 | 521 | |
| 522 | \medskip Our intended invariant is achieved as follows: | |
| 523 |   \begin{enumerate}
 | |
| 524 | ||
| 525 |   \item @{ML Wellformed_Terms.add} only admits terms that have passed
 | |
| 526 |   the @{ML Sign.cert_term} check of the given theory at that point.
 | |
| 527 | ||
| 528 |   \item Wellformedness in the sense of @{ML Sign.cert_term} is
 | |
| 529 | monotonic wrt.\ the sub-theory relation. So our data can move | |
| 530 | upwards in the hierarchy (via extension or merges), and maintain | |
| 531 | wellformedness without further checks. | |
| 532 | ||
| 533 |   \end{enumerate}
 | |
| 534 | ||
| 535 | Note that all basic operations of the inference kernel (which | |
| 536 |   includes @{ML Sign.cert_term}) observe this monotonicity principle,
 | |
| 537 | but other user-space tools don't. For example, fully-featured | |
| 538 |   type-inference via @{ML Syntax.check_term} (cf.\
 | |
| 539 |   \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
 | |
| 540 | background theory, since constraints of term constants can be | |
| 39821 | 541 | modified by later declarations, for example. | 
| 34928 | 542 | |
| 543 | In most cases, user-space context data does not have to take such | |
| 544 | invariants too seriously. The situation is different in the | |
| 545 | implementation of the inference kernel itself, which uses the very | |
| 546 | same data mechanisms for types, constants, axioms etc. | |
| 547 | *} | |
| 548 | ||
| 20447 | 549 | |
| 39865 | 550 | subsection {* Configuration options \label{sec:config-options} *}
 | 
| 551 | ||
| 552 | text {* A \emph{configuration option} is a named optional value of
 | |
| 553 | some basic type (Boolean, integer, string) that is stored in the | |
| 554 | context. It is a simple application of general context data | |
| 555 |   (\secref{sec:context-data}) that is sufficiently common to justify
 | |
| 556 | customized setup, which includes some concrete declarations for | |
| 557 | end-users using existing notation for attributes (cf.\ | |
| 558 |   \secref{sec:attributes}).
 | |
| 559 | ||
| 560 |   For example, the predefined configuration option @{attribute
 | |
| 561 | show_types} controls output of explicit type constraints for | |
| 39876 
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
 wenzelm parents: 
39866diff
changeset | 562 |   variables in printed terms (cf.\ \secref{sec:read-print}).  Its
 | 
| 39865 | 563 | value can be modified within Isar text like this: | 
| 564 | *} | |
| 565 | ||
| 566 | declare [[show_types = false]] | |
| 567 |   -- {* declaration within (local) theory context *}
 | |
| 568 | ||
| 40964 | 569 | notepad | 
| 570 | begin | |
| 39865 | 571 | note [[show_types = true]] | 
| 572 |     -- {* declaration within proof (forward mode) *}
 | |
| 573 | term x | |
| 574 | ||
| 575 | have "x = x" | |
| 576 | using [[show_types = false]] | |
| 577 |       -- {* declaration within proof (backward mode) *}
 | |
| 578 | .. | |
| 40964 | 579 | end | 
| 39865 | 580 | |
| 581 | text {* Configuration options that are not set explicitly hold a
 | |
| 582 | default value that can depend on the application context. This | |
| 583 | allows to retrieve the value from another slot within the context, | |
| 584 | or fall back on a global preference mechanism, for example. | |
| 585 | ||
| 586 | The operations to declare configuration options and get/map their | |
| 587 | values are modeled as direct replacements for historic global | |
| 588 | references, only that the context is made explicit. This allows | |
| 589 | easy configuration of tools, without relying on the execution order | |
| 590 | as required for old-style mutable references. *} | |
| 591 | ||
| 592 | text %mlref {*
 | |
| 593 |   \begin{mldecls}
 | |
| 594 |   @{index_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\
 | |
| 595 |   @{index_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\
 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 596 |   @{index_ML Attrib.setup_config_bool: "binding -> (Context.generic -> bool) ->
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 597 | bool Config.T"} \\ | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 598 |   @{index_ML Attrib.setup_config_int: "binding -> (Context.generic -> int) ->
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 599 | int Config.T"} \\ | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 600 |   @{index_ML Attrib.setup_config_real: "binding -> (Context.generic -> real) ->
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 601 | real Config.T"} \\ | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 602 |   @{index_ML Attrib.setup_config_string: "binding -> (Context.generic -> string) ->
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 603 | string Config.T"} \\ | 
| 39865 | 604 |   \end{mldecls}
 | 
| 605 | ||
| 606 |   \begin{description}
 | |
| 607 | ||
| 608 |   \item @{ML Config.get}~@{text "ctxt config"} gets the value of
 | |
| 609 |   @{text "config"} in the given context.
 | |
| 610 | ||
| 611 |   \item @{ML Config.map}~@{text "config f ctxt"} updates the context
 | |
| 612 |   by updating the value of @{text "config"}.
 | |
| 613 | ||
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 614 |   \item @{text "config ="}~@{ML Attrib.setup_config_bool}~@{text "name
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 615 |   default"} creates a named configuration option of type @{ML_type
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 616 |   bool}, with the given @{text "default"} depending on the application
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 617 |   context.  The resulting @{text "config"} can be used to get/map its
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 618 | value in a given context. There is an implicit update of the | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 619 | background theory that registers the option as attribute with some | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 620 | concrete syntax. | 
| 39865 | 621 | |
| 40291 | 622 |   \item @{ML Attrib.config_int}, @{ML Attrib.config_real}, and @{ML
 | 
| 623 |   Attrib.config_string} work like @{ML Attrib.config_bool}, but for
 | |
| 624 |   types @{ML_type int} and @{ML_type string}, respectively.
 | |
| 39865 | 625 | |
| 626 |   \end{description}
 | |
| 627 | *} | |
| 628 | ||
| 629 | text %mlex {* The following example shows how to declare and use a
 | |
| 630 |   Boolean configuration option called @{text "my_flag"} with constant
 | |
| 631 |   default value @{ML false}.  *}
 | |
| 632 | ||
| 633 | ML {*
 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 634 | val my_flag = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 635 |     Attrib.setup_config_bool @{binding my_flag} (K false)
 | 
| 39865 | 636 | *} | 
| 637 | ||
| 638 | text {* Now the user can refer to @{attribute my_flag} in
 | |
| 40126 | 639 | declarations, while ML tools can retrieve the current value from the | 
| 39865 | 640 |   context via @{ML Config.get}.  *}
 | 
| 641 | ||
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 642 | ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
 | 
| 39865 | 643 | |
| 644 | declare [[my_flag = true]] | |
| 645 | ||
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 646 | ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
 | 
| 39865 | 647 | |
| 40964 | 648 | notepad | 
| 649 | begin | |
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 650 |   {
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 651 | note [[my_flag = false]] | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 652 |     ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 653 | } | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 654 |   ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
 | 
| 40964 | 655 | end | 
| 39865 | 656 | |
| 40291 | 657 | text {* Here is another example involving ML type @{ML_type real}
 | 
| 658 | (floating-point numbers). *} | |
| 659 | ||
| 660 | ML {*
 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 661 | val airspeed_velocity = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 662 |     Attrib.setup_config_real @{binding airspeed_velocity} (K 0.0)
 | 
| 40291 | 663 | *} | 
| 664 | ||
| 40296 | 665 | declare [[airspeed_velocity = 10]] | 
| 40291 | 666 | declare [[airspeed_velocity = 9.9]] | 
| 667 | ||
| 39865 | 668 | |
| 26872 | 669 | section {* Names \label{sec:names} *}
 | 
| 20451 | 670 | |
| 34925 | 671 | text {* In principle, a name is just a string, but there are various
 | 
| 672 | conventions for representing additional structure. For example, | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 673 |   ``@{text "Foo.bar.baz"}'' is considered as a long name consisting of
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 674 |   qualifier @{text "Foo.bar"} and base name @{text "baz"}.  The
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 675 | individual constituents of a name may have further substructure, | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 676 | e.g.\ the string ``\verb,\,\verb,<alpha>,'' encodes as a single | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 677 | symbol. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 678 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 679 | \medskip Subsequently, we shall introduce specific categories of | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 680 | names. Roughly speaking these correspond to logical entities as | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 681 | follows: | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 682 |   \begin{itemize}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 683 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 684 |   \item Basic names (\secref{sec:basic-name}): free and bound
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 685 | variables. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 686 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 687 |   \item Indexed names (\secref{sec:indexname}): schematic variables.
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 688 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 689 |   \item Long names (\secref{sec:long-name}): constants of any kind
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 690 | (type constructors, term constants, other concepts defined in user | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 691 | space). Such entities are typically managed via name spaces | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 692 |   (\secref{sec:name-space}).
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 693 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 694 |   \end{itemize}
 | 
| 20451 | 695 | *} | 
| 20437 | 696 | |
| 697 | ||
| 39863 | 698 | subsection {* Strings of symbols \label{sec:symbols} *}
 | 
| 20437 | 699 | |
| 34925 | 700 | text {* A \emph{symbol} constitutes the smallest textual unit in
 | 
| 701 | Isabelle --- raw ML characters are normally not encountered at all! | |
| 702 | Isabelle strings consist of a sequence of symbols, represented as a | |
| 703 | packed string or an exploded list of strings. Each symbol is in | |
| 704 | itself a small string, which has either one of the following forms: | |
| 20437 | 705 | |
| 20451 | 706 |   \begin{enumerate}
 | 
| 20437 | 707 | |
| 37533 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 708 |   \item a single ASCII character ``@{text "c"}'', for example
 | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 709 | ``\verb,a,'', | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 710 | |
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 711 | \item a codepoint according to UTF8 (non-ASCII byte sequence), | 
| 20437 | 712 | |
| 20488 | 713 |   \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'',
 | 
| 20476 | 714 | for example ``\verb,\,\verb,<alpha>,'', | 
| 20437 | 715 | |
| 20488 | 716 |   \item a control symbol ``\verb,\,\verb,<^,@{text "ident"}\verb,>,'',
 | 
| 20476 | 717 | for example ``\verb,\,\verb,<^bold>,'', | 
| 20437 | 718 | |
| 20488 | 719 |   \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,''
 | 
| 34925 | 720 |   where @{text text} consists of printable characters excluding
 | 
| 20476 | 721 | ``\verb,.,'' and ``\verb,>,'', for example | 
| 722 |   ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
 | |
| 20437 | 723 | |
| 20488 | 724 |   \item a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text
 | 
| 20476 | 725 |   n}\verb,>, where @{text n} consists of digits, for example
 | 
| 20451 | 726 | ``\verb,\,\verb,<^raw42>,''. | 
| 20437 | 727 | |
| 20451 | 728 |   \end{enumerate}
 | 
| 20437 | 729 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 730 |   The @{text "ident"} syntax for symbol names is @{text "letter
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 731 |   (letter | digit)\<^sup>*"}, where @{text "letter = A..Za..z"} and @{text
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 732 | "digit = 0..9"}. There are infinitely many regular symbols and | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 733 | control symbols, but a fixed collection of standard symbols is | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 734 | treated specifically. For example, ``\verb,\,\verb,<alpha>,'' is | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 735 | classified as a letter, which means it may occur within regular | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 736 | Isabelle identifiers. | 
| 20437 | 737 | |
| 37533 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 738 | The character set underlying Isabelle symbols is 7-bit ASCII, but | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 739 | 8-bit character sequences are passed-through unchanged. Unicode/UCS | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 740 | data in UTF-8 encoding is processed in a non-strict fashion, such | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 741 | that well-formed code sequences are recognized | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 742 |   accordingly.\footnote{Note that ISO-Latin-1 differs from UTF-8 only
 | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 743 | in some special punctuation characters that even have replacements | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 744 | within the standard collection of Isabelle symbols. Text consisting | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 745 | of ASCII plus accented letters can be processed in either encoding.} | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 746 | Unicode provides its own collection of mathematical symbols, but | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 747 | within the core Isabelle/ML world there is no link to the standard | 
| 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 wenzelm parents: 
36611diff
changeset | 748 | collection of Isabelle regular symbols. | 
| 20476 | 749 | |
| 750 | \medskip Output of Isabelle symbols depends on the print mode | |
| 29758 | 751 |   (\secref{print-mode}).  For example, the standard {\LaTeX} setup of
 | 
| 752 | the Isabelle document preparation system would present | |
| 20451 | 753 |   ``\verb,\,\verb,<alpha>,'' as @{text "\<alpha>"}, and
 | 
| 754 |   ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text
 | |
| 34925 | 755 | "\<^bold>\<alpha>"}. On-screen rendering usually works by mapping a finite | 
| 756 | subset of Isabelle symbols to suitable Unicode characters. | |
| 20451 | 757 | *} | 
| 20437 | 758 | |
| 759 | text %mlref {*
 | |
| 760 |   \begin{mldecls}
 | |
| 34921 | 761 |   @{index_ML_type "Symbol.symbol": string} \\
 | 
| 20437 | 762 |   @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
 | 
| 763 |   @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
 | |
| 764 |   @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
 | |
| 765 |   @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
 | |
| 20547 | 766 |   @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
 | 
| 767 |   \end{mldecls}
 | |
| 768 |   \begin{mldecls}
 | |
| 20437 | 769 |   @{index_ML_type "Symbol.sym"} \\
 | 
| 770 |   @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
 | |
| 771 |   \end{mldecls}
 | |
| 772 | ||
| 773 |   \begin{description}
 | |
| 774 | ||
| 39864 | 775 |   \item Type @{ML_type "Symbol.symbol"} represents individual Isabelle
 | 
| 34921 | 776 | symbols. | 
| 20437 | 777 | |
| 20476 | 778 |   \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list
 | 
| 39821 | 779 |   from the packed form.  This function supersedes @{ML
 | 
| 20476 | 780 | "String.explode"} for virtually all purposes of manipulating text in | 
| 34925 | 781 |   Isabelle!\footnote{The runtime overhead for exploded strings is
 | 
| 782 | mainly that of the list structure: individual symbols that happen to | |
| 39821 | 783 | be a singleton string do not require extra memory in Poly/ML.} | 
| 20437 | 784 | |
| 785 |   \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
 | |
| 20476 | 786 |   "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
 | 
| 787 | symbols according to fixed syntactic conventions of Isabelle, cf.\ | |
| 788 |   \cite{isabelle-isar-ref}.
 | |
| 20437 | 789 | |
| 39864 | 790 |   \item Type @{ML_type "Symbol.sym"} is a concrete datatype that
 | 
| 791 | represents the different kinds of symbols explicitly, with | |
| 792 |   constructors @{ML "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML
 | |
| 793 |   "Symbol.UTF8"}, @{ML "Symbol.Ctrl"}, @{ML "Symbol.Raw"}.
 | |
| 20437 | 794 | |
| 795 |   \item @{ML "Symbol.decode"} converts the string representation of a
 | |
| 20451 | 796 | symbol into the datatype version. | 
| 20437 | 797 | |
| 798 |   \end{description}
 | |
| 34925 | 799 | |
| 800 |   \paragraph{Historical note.} In the original SML90 standard the
 | |
| 40628 | 801 |   primitive ML type @{ML_type char} did not exists, and the @{ML_text
 | 
| 34925 | 802 | "explode: string -> string list"} operation would produce a list of | 
| 40628 | 803 |   singleton strings as does @{ML "raw_explode: string -> string list"}
 | 
| 804 | in Isabelle/ML today. When SML97 came out, Isabelle did not adopt | |
| 805 | its slightly anachronistic 8-bit characters, but the idea of | |
| 806 | exploding a string into a list of small strings was extended to | |
| 807 | ``symbols'' as explained above. Thus Isabelle sources can refer to | |
| 808 | an infinite store of user-defined symbols, without having to worry | |
| 809 | about the multitude of Unicode encodings. *} | |
| 20437 | 810 | |
| 811 | ||
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 812 | subsection {* Basic names \label{sec:basic-name} *}
 | 
| 20476 | 813 | |
| 814 | text {*
 | |
| 815 |   A \emph{basic name} essentially consists of a single Isabelle
 | |
| 816 | identifier. There are conventions to mark separate classes of basic | |
| 29761 | 817 | names, by attaching a suffix of underscores: one underscore means | 
| 818 |   \emph{internal name}, two underscores means \emph{Skolem name},
 | |
| 819 |   three underscores means \emph{internal Skolem name}.
 | |
| 20476 | 820 | |
| 821 |   For example, the basic name @{text "foo"} has the internal version
 | |
| 822 |   @{text "foo_"}, with Skolem versions @{text "foo__"} and @{text
 | |
| 823 | "foo___"}, respectively. | |
| 824 | ||
| 20488 | 825 | These special versions provide copies of the basic name space, apart | 
| 826 | from anything that normally appears in the user text. For example, | |
| 827 | system generated variables in Isar proof contexts are usually marked | |
| 34926 | 828 |   as internal, which prevents mysterious names like @{text "xaa"} to
 | 
| 829 | appear in human-readable text. | |
| 20476 | 830 | |
| 20488 | 831 | \medskip Manipulating binding scopes often requires on-the-fly | 
| 832 |   renamings.  A \emph{name context} contains a collection of already
 | |
| 833 |   used names.  The @{text "declare"} operation adds names to the
 | |
| 834 | context. | |
| 20476 | 835 | |
| 20488 | 836 |   The @{text "invents"} operation derives a number of fresh names from
 | 
| 837 | a given starting point. For example, the first three names derived | |
| 838 |   from @{text "a"} are @{text "a"}, @{text "b"}, @{text "c"}.
 | |
| 20476 | 839 | |
| 840 |   The @{text "variants"} operation produces fresh names by
 | |
| 20488 | 841 |   incrementing tentative names as base-26 numbers (with digits @{text
 | 
| 842 |   "a..z"}) until all clashes are resolved.  For example, name @{text
 | |
| 843 |   "foo"} results in variants @{text "fooa"}, @{text "foob"}, @{text
 | |
| 844 |   "fooc"}, \dots, @{text "fooaa"}, @{text "fooab"} etc.; each renaming
 | |
| 845 | step picks the next unused variant from this sequence. | |
| 20476 | 846 | *} | 
| 847 | ||
| 848 | text %mlref {*
 | |
| 849 |   \begin{mldecls}
 | |
| 850 |   @{index_ML Name.internal: "string -> string"} \\
 | |
| 20547 | 851 |   @{index_ML Name.skolem: "string -> string"} \\
 | 
| 852 |   \end{mldecls}
 | |
| 853 |   \begin{mldecls}
 | |
| 20476 | 854 |   @{index_ML_type Name.context} \\
 | 
| 855 |   @{index_ML Name.context: Name.context} \\
 | |
| 856 |   @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
 | |
| 857 |   @{index_ML Name.invents: "Name.context -> string -> int -> string list"} \\
 | |
| 858 |   @{index_ML Name.variants: "string list -> Name.context -> string list * Name.context"} \\
 | |
| 859 |   \end{mldecls}
 | |
| 34926 | 860 |   \begin{mldecls}
 | 
| 861 |   @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
 | |
| 862 |   \end{mldecls}
 | |
| 20476 | 863 | |
| 864 |   \begin{description}
 | |
| 865 | ||
| 866 |   \item @{ML Name.internal}~@{text "name"} produces an internal name
 | |
| 867 | by adding one underscore. | |
| 868 | ||
| 869 |   \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by
 | |
| 870 | adding two underscores. | |
| 871 | ||
| 39864 | 872 |   \item Type @{ML_type Name.context} represents the context of already
 | 
| 873 |   used names; the initial value is @{ML "Name.context"}.
 | |
| 20476 | 874 | |
| 20488 | 875 |   \item @{ML Name.declare}~@{text "name"} enters a used name into the
 | 
| 876 | context. | |
| 20437 | 877 | |
| 20488 | 878 |   \item @{ML Name.invents}~@{text "context name n"} produces @{text
 | 
| 879 |   "n"} fresh names derived from @{text "name"}.
 | |
| 880 | ||
| 881 |   \item @{ML Name.variants}~@{text "names context"} produces fresh
 | |
| 29761 | 882 |   variants of @{text "names"}; the result is entered into the context.
 | 
| 20476 | 883 | |
| 34926 | 884 |   \item @{ML Variable.names_of}~@{text "ctxt"} retrieves the context
 | 
| 885 | of declared type and term variable names. Projecting a proof | |
| 886 | context down to a primitive name context is occasionally useful when | |
| 887 | invoking lower-level operations. Regular management of ``fresh | |
| 888 |   variables'' is done by suitable operations of structure @{ML_struct
 | |
| 889 | Variable}, which is also able to provide an official status of | |
| 890 | ``locally fixed variable'' within the logical environment (cf.\ | |
| 891 |   \secref{sec:variables}).
 | |
| 892 | ||
| 20476 | 893 |   \end{description}
 | 
| 894 | *} | |
| 895 | ||
| 39857 | 896 | text %mlex {* The following simple examples demonstrate how to produce
 | 
| 897 |   fresh names from the initial @{ML Name.context}. *}
 | |
| 898 | ||
| 899 | ML {*
 | |
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 900 | val list1 = Name.invents Name.context "a" 5; | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 901 |   @{assert} (list1 = ["a", "b", "c", "d", "e"]);
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 902 | |
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 903 | val list2 = | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 904 | #1 (Name.variants ["x", "x", "a", "a", "'a", "'a"] Name.context); | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 905 |   @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]);
 | 
| 39857 | 906 | *} | 
| 907 | ||
| 40126 | 908 | text {* \medskip The same works relatively to the formal context as
 | 
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 909 | follows. *} | 
| 39857 | 910 | |
| 911 | locale ex = fixes a b c :: 'a | |
| 912 | begin | |
| 913 | ||
| 914 | ML {*
 | |
| 915 |   val names = Variable.names_of @{context};
 | |
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 916 | |
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 917 | val list1 = Name.invents names "a" 5; | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 918 |   @{assert} (list1 = ["d", "e", "f", "g", "h"]);
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 919 | |
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 920 | val list2 = | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 921 | #1 (Name.variants ["x", "x", "a", "a", "'a", "'a"] names); | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 922 |   @{assert} (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]);
 | 
| 39857 | 923 | *} | 
| 924 | ||
| 925 | end | |
| 926 | ||
| 20476 | 927 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 928 | subsection {* Indexed names \label{sec:indexname} *}
 | 
| 20476 | 929 | |
| 930 | text {*
 | |
| 931 |   An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic
 | |
| 20488 | 932 | name and a natural number. This representation allows efficient | 
| 933 | renaming by incrementing the second component only. The canonical | |
| 934 | way to rename two collections of indexnames apart from each other is | |
| 935 |   this: determine the maximum index @{text "maxidx"} of the first
 | |
| 936 | collection, then increment all indexes of the second collection by | |
| 937 |   @{text "maxidx + 1"}; the maximum index of an empty collection is
 | |
| 938 |   @{text "-1"}.
 | |
| 20476 | 939 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 940 | Occasionally, basic names are injected into the same pair type of | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 941 |   indexed names: then @{text "(x, -1)"} is used to encode the basic
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 942 |   name @{text "x"}.
 | 
| 20488 | 943 | |
| 944 | \medskip Isabelle syntax observes the following rules for | |
| 945 |   representing an indexname @{text "(x, i)"} as a packed string:
 | |
| 20476 | 946 | |
| 947 |   \begin{itemize}
 | |
| 948 | ||
| 20479 | 949 |   \item @{text "?x"} if @{text "x"} does not end with a digit and @{text "i = 0"},
 | 
| 20476 | 950 | |
| 951 |   \item @{text "?xi"} if @{text "x"} does not end with a digit,
 | |
| 952 | ||
| 20488 | 953 |   \item @{text "?x.i"} otherwise.
 | 
| 20476 | 954 | |
| 955 |   \end{itemize}
 | |
| 20470 | 956 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 957 | Indexnames may acquire large index numbers after several maxidx | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 958 | shifts have been applied. Results are usually normalized towards | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 959 |   @{text "0"} at certain checkpoints, notably at the end of a proof.
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 960 | This works by producing variants of the corresponding basic name | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 961 |   components.  For example, the collection @{text "?x1, ?x7, ?x42"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 962 |   becomes @{text "?x, ?xa, ?xb"}.
 | 
| 20476 | 963 | *} | 
| 964 | ||
| 965 | text %mlref {*
 | |
| 966 |   \begin{mldecls}
 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 967 |   @{index_ML_type indexname: "string * int"} \\
 | 
| 20476 | 968 |   \end{mldecls}
 | 
| 969 | ||
| 970 |   \begin{description}
 | |
| 971 | ||
| 39864 | 972 |   \item Type @{ML_type indexname} represents indexed names.  This is
 | 
| 973 |   an abbreviation for @{ML_type "string * int"}.  The second component
 | |
| 974 |   is usually non-negative, except for situations where @{text "(x,
 | |
| 975 | -1)"} is used to inject basic names into this type. Other negative | |
| 34926 | 976 | indexes should not be used. | 
| 20476 | 977 | |
| 978 |   \end{description}
 | |
| 979 | *} | |
| 980 | ||
| 981 | ||
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 982 | subsection {* Long names \label{sec:long-name} *}
 | 
| 20476 | 983 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 984 | text {* A \emph{long name} consists of a sequence of non-empty name
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 985 | components. The packed representation uses a dot as separator, as | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 986 |   in ``@{text "A.b.c"}''.  The last component is called \emph{base
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 987 |   name}, the remaining prefix is called \emph{qualifier} (which may be
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 988 | empty). The qualifier can be understood as the access path to the | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 989 | named entity while passing through some nested block-structure, | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 990 | although our free-form long names do not really enforce any strict | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 991 | discipline. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 992 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 993 |   For example, an item named ``@{text "A.b.c"}'' may be understood as
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 994 |   a local entity @{text "c"}, within a local structure @{text "b"},
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 995 |   within a global structure @{text "A"}.  In practice, long names
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 996 | usually represent 1--3 levels of qualification. User ML code should | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 997 | not make any assumptions about the particular structure of long | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 998 | names! | 
| 20437 | 999 | |
| 20476 | 1000 | The empty name is commonly used as an indication of unnamed | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1001 | entities, or entities that are not entered into the corresponding | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1002 | name space, whenever this makes any sense. The basic operations on | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1003 | long names map empty names again to empty names. | 
| 20437 | 1004 | *} | 
| 1005 | ||
| 20476 | 1006 | text %mlref {*
 | 
| 1007 |   \begin{mldecls}
 | |
| 30365 | 1008 |   @{index_ML Long_Name.base_name: "string -> string"} \\
 | 
| 1009 |   @{index_ML Long_Name.qualifier: "string -> string"} \\
 | |
| 1010 |   @{index_ML Long_Name.append: "string -> string -> string"} \\
 | |
| 1011 |   @{index_ML Long_Name.implode: "string list -> string"} \\
 | |
| 1012 |   @{index_ML Long_Name.explode: "string -> string list"} \\
 | |
| 20547 | 1013 |   \end{mldecls}
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1014 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1015 |   \begin{description}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1016 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1017 |   \item @{ML Long_Name.base_name}~@{text "name"} returns the base name
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1018 | of a long name. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1019 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1020 |   \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1021 | of a long name. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1022 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1023 |   \item @{ML Long_Name.append}~@{text "name\<^isub>1 name\<^isub>2"} appends two long
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1024 | names. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1025 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1026 |   \item @{ML Long_Name.implode}~@{text "names"} and @{ML
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1027 |   Long_Name.explode}~@{text "name"} convert between the packed string
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1028 | representation and the explicit list form of long names. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1029 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1030 |   \end{description}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1031 | *} | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1032 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1033 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1034 | subsection {* Name spaces \label{sec:name-space} *}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1035 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1036 | text {* A @{text "name space"} manages a collection of long names,
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1037 | together with a mapping between partially qualified external names | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1038 | and fully qualified internal names (in both directions). Note that | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1039 |   the corresponding @{text "intern"} and @{text "extern"} operations
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1040 |   are mostly used for parsing and printing only!  The @{text
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1041 | "declare"} operation augments a name space according to the accesses | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1042 | determined by a given binding, and a naming policy from the context. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1043 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1044 |   \medskip A @{text "binding"} specifies details about the prospective
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1045 | long name of a newly introduced formal entity. It consists of a | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1046 | base name, prefixes for qualification (separate ones for system | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1047 | infrastructure and user-space mechanisms), a slot for the original | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1048 | source position, and some additional flags. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1049 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1050 |   \medskip A @{text "naming"} provides some additional details for
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1051 | producing a long name from a binding. Normally, the naming is | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1052 |   implicit in the theory or proof context.  The @{text "full"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1053 | operation (and its variants for different context types) produces a | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1054 | fully qualified internal name to be entered into a name space. The | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1055 | main equation of this ``chemical reaction'' when binding new | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1056 | entities in a context is as follows: | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1057 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1058 | \medskip | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1059 |   \begin{tabular}{l}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1060 |   @{text "binding + naming \<longrightarrow> long name + name space accesses"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1061 |   \end{tabular}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1062 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1063 | \bigskip As a general principle, there is a separate name space for | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1064 | each kind of formal entity, e.g.\ fact, logical constant, type | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1065 | constructor, type class. It is usually clear from the occurrence in | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1066 | concrete syntax (or from the scope) which kind of entity a name | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1067 |   refers to.  For example, the very same name @{text "c"} may be used
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1068 | uniformly for a constant, type constructor, and type class. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1069 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1070 | There are common schemes to name derived entities systematically | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1071 | according to the name of the main logical entity involved, e.g.\ | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1072 |   fact @{text "c.intro"} for a canonical introduction rule related to
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1073 |   constant @{text "c"}.  This technique of mapping names from one
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1074 | space into another requires some care in order to avoid conflicts. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1075 | In particular, theorem names derived from a type constructor or type | 
| 39839 | 1076 | class should get an additional suffix in addition to the usual | 
| 1077 | qualification. This leads to the following conventions for derived | |
| 1078 | names: | |
| 1079 | ||
| 1080 | \medskip | |
| 1081 |   \begin{tabular}{ll}
 | |
| 1082 | logical entity & fact name \\\hline | |
| 1083 |   constant @{text "c"} & @{text "c.intro"} \\
 | |
| 1084 |   type @{text "c"} & @{text "c_type.intro"} \\
 | |
| 1085 |   class @{text "c"} & @{text "c_class.intro"} \\
 | |
| 1086 |   \end{tabular}
 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1087 | *} | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1088 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1089 | text %mlref {*
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1090 |   \begin{mldecls}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1091 |   @{index_ML_type binding} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1092 |   @{index_ML Binding.empty: binding} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1093 |   @{index_ML Binding.name: "string -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1094 |   @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1095 |   @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1096 |   @{index_ML Binding.conceal: "binding -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1097 |   @{index_ML Binding.str_of: "binding -> string"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1098 |   \end{mldecls}
 | 
| 20547 | 1099 |   \begin{mldecls}
 | 
| 33174 | 1100 |   @{index_ML_type Name_Space.naming} \\
 | 
| 1101 |   @{index_ML Name_Space.default_naming: Name_Space.naming} \\
 | |
| 1102 |   @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
 | |
| 1103 |   @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
 | |
| 20547 | 1104 |   \end{mldecls}
 | 
| 1105 |   \begin{mldecls}
 | |
| 33174 | 1106 |   @{index_ML_type Name_Space.T} \\
 | 
| 1107 |   @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
 | |
| 1108 |   @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
 | |
| 42401 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 wenzelm parents: 
42361diff
changeset | 1109 |   @{index_ML Name_Space.declare: "Proof.context -> bool ->
 | 
| 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 wenzelm parents: 
42361diff
changeset | 1110 | Name_Space.naming -> binding -> Name_Space.T -> string * Name_Space.T"} \\ | 
| 33174 | 1111 |   @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
 | 
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
40964diff
changeset | 1112 |   @{index_ML Name_Space.extern: "Proof.context -> Name_Space.T -> string -> string"} \\
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1113 |   @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
 | 
| 20476 | 1114 |   \end{mldecls}
 | 
| 20437 | 1115 | |
| 20476 | 1116 |   \begin{description}
 | 
| 1117 | ||
| 39864 | 1118 |   \item Type @{ML_type binding} represents the abstract concept of
 | 
| 1119 | name bindings. | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1120 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1121 |   \item @{ML Binding.empty} is the empty binding.
 | 
| 20476 | 1122 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1123 |   \item @{ML Binding.name}~@{text "name"} produces a binding with base
 | 
| 39832 | 1124 |   name @{text "name"}.  Note that this lacks proper source position
 | 
| 1125 |   information; see also the ML antiquotation @{ML_antiquotation
 | |
| 1126 | binding}. | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1127 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1128 |   \item @{ML Binding.qualify}~@{text "mandatory name binding"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1129 |   prefixes qualifier @{text "name"} to @{text "binding"}.  The @{text
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1130 | "mandatory"} flag tells if this name component always needs to be | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1131 |   given in name space accesses --- this is mostly @{text "false"} in
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1132 | practice. Note that this part of qualification is typically used in | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1133 | derived specification mechanisms. | 
| 20437 | 1134 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1135 |   \item @{ML Binding.prefix} is similar to @{ML Binding.qualify}, but
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1136 | affects the system prefix. This part of extra qualification is | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1137 | typically used in the infrastructure for modular specifications, | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1138 |   notably ``local theory targets'' (see also \chref{ch:local-theory}).
 | 
| 20437 | 1139 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1140 |   \item @{ML Binding.conceal}~@{text "binding"} indicates that the
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1141 | binding shall refer to an entity that serves foundational purposes | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1142 | only. This flag helps to mark implementation details of | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1143 | specification mechanism etc. Other tools should not depend on the | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1144 |   particulars of concealed entities (cf.\ @{ML
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1145 | Name_Space.is_concealed}). | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1146 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1147 |   \item @{ML Binding.str_of}~@{text "binding"} produces a string
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1148 | representation for human-readable output, together with some formal | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1149 | markup that might get used in GUI front-ends, for example. | 
| 20476 | 1150 | |
| 39864 | 1151 |   \item Type @{ML_type Name_Space.naming} represents the abstract
 | 
| 1152 | concept of a naming policy. | |
| 20437 | 1153 | |
| 33174 | 1154 |   \item @{ML Name_Space.default_naming} is the default naming policy.
 | 
| 20476 | 1155 | In a theory context, this is usually augmented by a path prefix | 
| 1156 | consisting of the theory name. | |
| 1157 | ||
| 33174 | 1158 |   \item @{ML Name_Space.add_path}~@{text "path naming"} augments the
 | 
| 20488 | 1159 | naming policy by extending its path component. | 
| 20437 | 1160 | |
| 33174 | 1161 |   \item @{ML Name_Space.full_name}~@{text "naming binding"} turns a
 | 
| 30281 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 wenzelm parents: 
30272diff
changeset | 1162 | name binding (usually a basic name) into the fully qualified | 
| 29008 | 1163 | internal name, according to the given naming policy. | 
| 20476 | 1164 | |
| 39864 | 1165 |   \item Type @{ML_type Name_Space.T} represents name spaces.
 | 
| 20476 | 1166 | |
| 33174 | 1167 |   \item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
 | 
| 20488 | 1168 | "(space\<^isub>1, space\<^isub>2)"} are the canonical operations for | 
| 1169 | maintaining name spaces according to theory data management | |
| 33174 | 1170 |   (\secref{sec:context-data}); @{text "kind"} is a formal comment
 | 
| 1171 | to characterize the purpose of a name space. | |
| 20437 | 1172 | |
| 42401 
9bfaf6819291
updated some theory primitives, which now depend on auxiliary context;
 wenzelm parents: 
42361diff
changeset | 1173 |   \item @{ML Name_Space.declare}~@{text "ctxt strict naming bindings
 | 
| 33174 | 1174 | space"} enters a name binding as fully qualified internal name into | 
| 1175 | the name space, with external accesses determined by the naming | |
| 1176 | policy. | |
| 20476 | 1177 | |
| 33174 | 1178 |   \item @{ML Name_Space.intern}~@{text "space name"} internalizes a
 | 
| 20476 | 1179 | (partially qualified) external name. | 
| 20437 | 1180 | |
| 20488 | 1181 | This operation is mostly for parsing! Note that fully qualified | 
| 20476 | 1182 |   names stemming from declarations are produced via @{ML
 | 
| 33174 | 1183 |   "Name_Space.full_name"} and @{ML "Name_Space.declare"}
 | 
| 29008 | 1184 |   (or their derivatives for @{ML_type theory} and
 | 
| 20488 | 1185 |   @{ML_type Proof.context}).
 | 
| 20437 | 1186 | |
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
40964diff
changeset | 1187 |   \item @{ML Name_Space.extern}~@{text "ctxt space name"} externalizes a
 | 
| 20476 | 1188 | (fully qualified) internal name. | 
| 1189 | ||
| 30281 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 wenzelm parents: 
30272diff
changeset | 1190 | This operation is mostly for printing! User code should not rely on | 
| 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 wenzelm parents: 
30272diff
changeset | 1191 | the precise result too much. | 
| 20476 | 1192 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1193 |   \item @{ML Name_Space.is_concealed}~@{text "space name"} indicates
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1194 |   whether @{text "name"} refers to a strictly private entity that
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1195 | other tools are supposed to ignore! | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1196 | |
| 20476 | 1197 |   \end{description}
 | 
| 1198 | *} | |
| 30272 | 1199 | |
| 39832 | 1200 | text %mlantiq {*
 | 
| 1201 |   \begin{matharray}{rcl}
 | |
| 1202 |   @{ML_antiquotation_def "binding"} & : & @{text ML_antiquotation} \\
 | |
| 1203 |   \end{matharray}
 | |
| 1204 | ||
| 42510 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 1205 |   @{rail "
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 1206 |   @@{ML_antiquotation binding} name
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 1207 | "} | 
| 39832 | 1208 | |
| 1209 |   \begin{description}
 | |
| 1210 | ||
| 1211 |   \item @{text "@{binding name}"} produces a binding with base name
 | |
| 1212 |   @{text "name"} and the source position taken from the concrete
 | |
| 1213 | syntax of this antiquotation. In many situations this is more | |
| 1214 |   appropriate than the more basic @{ML Binding.name} function.
 | |
| 1215 | ||
| 1216 |   \end{description}
 | |
| 1217 | *} | |
| 1218 | ||
| 39833 | 1219 | text %mlex {* The following example yields the source position of some
 | 
| 40126 | 1220 | concrete binding inlined into the text: | 
| 39833 | 1221 | *} | 
| 1222 | ||
| 1223 | ML {* Binding.pos_of @{binding here} *}
 | |
| 1224 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1225 | text {* \medskip That position can be also printed in a message as
 | 
| 40126 | 1226 | follows: *} | 
| 39833 | 1227 | |
| 1228 | ML_command {*
 | |
| 1229 | writeln | |
| 1230 |     ("Look here" ^ Position.str_of (Binding.pos_of @{binding here}))
 | |
| 1231 | *} | |
| 1232 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1233 | text {* This illustrates a key virtue of formalized bindings as
 | 
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1234 | opposed to raw specifications of base names: the system can use this | 
| 40126 | 1235 | additional information for feedback given to the user (error | 
| 1236 | messages etc.). *} | |
| 39833 | 1237 | |
| 18537 | 1238 | end |