| author | traytel | 
| Fri, 29 Nov 2013 08:26:45 +0100 | |
| changeset 54611 | 31afce809794 | 
| parent 53015 | a1119cf551e8 | 
| child 55112 | b1a5d603fd12 | 
| 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 |   The @{text "merge"} operation produces the least upper bound of two
 | 
| 84 | theories, which actually degenerates into absorption of one theory | |
| 85 | into the other (according to the nominal sub-theory relation). | |
| 86 | ||
| 87 |   The @{text "begin"} operation starts a new theory by importing
 | |
| 88 | several parent theories and entering a special mode of nameless | |
| 89 |   incremental updates, until the final @{text "end"} operation is
 | |
| 90 | performed. | |
| 91 | ||
| 20447 | 92 |   \medskip The example in \figref{fig:ex-theory} below shows a theory
 | 
| 20451 | 93 |   graph derived from @{text "Pure"}, with theory @{text "Length"}
 | 
| 94 |   importing @{text "Nat"} and @{text "List"}.  The body of @{text
 | |
| 52788 | 95 | "Length"} consists of a sequence of updates, resulting in locally a | 
| 96 | linear sub-theory relation for each intermediate step. | |
| 20447 | 97 | |
| 98 |   \begin{figure}[htb]
 | |
| 99 |   \begin{center}
 | |
| 20429 | 100 |   \begin{tabular}{rcccl}
 | 
| 20447 | 101 |         &            & @{text "Pure"} \\
 | 
| 102 |         &            & @{text "\<down>"} \\
 | |
| 103 |         &            & @{text "FOL"} \\
 | |
| 18537 | 104 | & $\swarrow$ & & $\searrow$ & \\ | 
| 21852 | 105 |   @{text "Nat"} &    &              &            & @{text "List"} \\
 | 
| 18537 | 106 | & $\searrow$ & & $\swarrow$ \\ | 
| 20447 | 107 |         &            & @{text "Length"} \\
 | 
| 26864 | 108 |         &            & \multicolumn{3}{l}{~~@{keyword "begin"}} \\
 | 
| 18537 | 109 | & & $\vdots$~~ \\ | 
| 26864 | 110 |         &            & \multicolumn{3}{l}{~~@{command "end"}} \\
 | 
| 20429 | 111 |   \end{tabular}
 | 
| 20451 | 112 |   \caption{A theory definition depending on ancestors}\label{fig:ex-theory}
 | 
| 20447 | 113 |   \end{center}
 | 
| 114 |   \end{figure}
 | |
| 20451 | 115 | |
| 52788 | 116 | \medskip Derived formal entities may retain a reference to the | 
| 117 | background theory in order to indicate the formal context from which | |
| 118 | they were produced. This provides an immutable certificate of the | |
| 119 | background theory. *} | |
| 18537 | 120 | |
| 20430 | 121 | text %mlref {*
 | 
| 20447 | 122 |   \begin{mldecls}
 | 
| 123 |   @{index_ML_type theory} \\
 | |
| 39837 | 124 |   @{index_ML Theory.eq_thy: "theory * theory -> bool"} \\
 | 
| 20447 | 125 |   @{index_ML Theory.subthy: "theory * theory -> bool"} \\
 | 
| 34921 | 126 |   @{index_ML Theory.merge: "theory * theory -> theory"} \\
 | 
| 48930 | 127 |   @{index_ML Theory.begin_theory: "string * Position.T -> theory list -> theory"} \\
 | 
| 39837 | 128 |   @{index_ML Theory.parents_of: "theory -> theory list"} \\
 | 
| 129 |   @{index_ML Theory.ancestors_of: "theory -> theory list"} \\
 | |
| 20547 | 130 |   \end{mldecls}
 | 
| 20447 | 131 | |
| 132 |   \begin{description}
 | |
| 133 | ||
| 52788 | 134 |   \item Type @{ML_type theory} represents theory contexts.
 | 
| 20447 | 135 | |
| 39837 | 136 |   \item @{ML "Theory.eq_thy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} check strict
 | 
| 137 | identity of two theories. | |
| 138 | ||
| 34921 | 139 |   \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} compares theories
 | 
| 140 | according to the intrinsic graph structure of the construction. | |
| 141 | This sub-theory relation is a nominal approximation of inclusion | |
| 142 |   (@{text "\<subseteq>"}) of the corresponding content (according to the
 | |
| 143 | semantics of the ML modules that implement the data). | |
| 20447 | 144 | |
| 34921 | 145 |   \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} absorbs one theory
 | 
| 52788 | 146 | into the other. This version of ad-hoc theory merge fails for | 
| 147 | unrelated theories! | |
| 34921 | 148 | |
| 149 |   \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 | 150 | a new theory based on the given parents. This ML function is | 
| 34921 | 151 | normally not invoked directly. | 
| 20447 | 152 | |
| 39837 | 153 |   \item @{ML "Theory.parents_of"}~@{text "thy"} returns the direct
 | 
| 154 |   ancestors of @{text thy}.
 | |
| 155 | ||
| 156 |   \item @{ML "Theory.ancestors_of"}~@{text "thy"} returns all
 | |
| 157 |   ancestors of @{text thy} (not including @{text thy} itself).
 | |
| 158 | ||
| 20447 | 159 |   \end{description}
 | 
| 20430 | 160 | *} | 
| 161 | ||
| 39832 | 162 | text %mlantiq {*
 | 
| 163 |   \begin{matharray}{rcl}
 | |
| 164 |   @{ML_antiquotation_def "theory"} & : & @{text ML_antiquotation} \\
 | |
| 51686 | 165 |   @{ML_antiquotation_def "theory_context"} & : & @{text ML_antiquotation} \\
 | 
| 39832 | 166 |   \end{matharray}
 | 
| 167 | ||
| 42510 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 168 |   @{rail "
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 169 |   @@{ML_antiquotation theory} nameref?
 | 
| 51686 | 170 | ; | 
| 171 |   @@{ML_antiquotation theory_context} nameref
 | |
| 42510 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 172 | "} | 
| 39832 | 173 | |
| 174 |   \begin{description}
 | |
| 175 | ||
| 176 |   \item @{text "@{theory}"} refers to the background theory of the
 | |
| 177 | current context --- as abstract value. | |
| 178 | ||
| 179 |   \item @{text "@{theory A}"} refers to an explicitly named ancestor
 | |
| 180 |   theory @{text "A"} of the background theory of the current context
 | |
| 181 | --- as abstract value. | |
| 182 | ||
| 51686 | 183 |   \item @{text "@{theory_context A}"} is similar to @{text "@{theory
 | 
| 184 |   A}"}, but presents the result as initial @{ML_type Proof.context}
 | |
| 185 |   (see also @{ML Proof_Context.init_global}).
 | |
| 186 | ||
| 39832 | 187 |   \end{description}
 | 
| 188 | *} | |
| 189 | ||
| 18537 | 190 | |
| 191 | subsection {* Proof context \label{sec:context-proof} *}
 | |
| 192 | ||
| 52788 | 193 | text {* A proof context is a container for pure data that refers to
 | 
| 194 |   the theory from which it is derived. The @{text "init"} operation
 | |
| 195 | creates a proof context from a given theory. There is an explicit | |
| 196 |   @{text "transfer"} operation to force resynchronization with updates
 | |
| 197 | to the background theory -- this is rarely required in practice. | |
| 20429 | 198 | |
| 34921 | 199 | Entities derived in a proof context need to record logical | 
| 20447 | 200 | requirements explicitly, since there is no separate context | 
| 34921 | 201 | identification or symbolic inclusion as for theories. For example, | 
| 202 |   hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
 | |
| 203 |   are recorded separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, just to
 | |
| 204 | make double sure. Results could still leak into an alien proof | |
| 205 | context due to programming errors, but Isabelle/Isar includes some | |
| 206 | extra validity checks in critical positions, notably at the end of a | |
| 207 | sub-proof. | |
| 20429 | 208 | |
| 20451 | 209 | Proof contexts may be manipulated arbitrarily, although the common | 
| 210 | discipline is to follow block structure as a mental model: a given | |
| 211 | context is extended consecutively, and results are exported back | |
| 34921 | 212 | into the original context. Note that an Isar proof state models | 
| 20451 | 213 | block-structured reasoning explicitly, using a stack of proof | 
| 34921 | 214 | contexts internally. For various technical reasons, the background | 
| 215 | theory of an Isar proof state must not be changed while the proof is | |
| 216 | still under construction! | |
| 18537 | 217 | *} | 
| 218 | ||
| 20449 | 219 | text %mlref {*
 | 
| 220 |   \begin{mldecls}
 | |
| 221 |   @{index_ML_type Proof.context} \\
 | |
| 42361 | 222 |   @{index_ML Proof_Context.init_global: "theory -> Proof.context"} \\
 | 
| 223 |   @{index_ML Proof_Context.theory_of: "Proof.context -> theory"} \\
 | |
| 224 |   @{index_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\
 | |
| 20449 | 225 |   \end{mldecls}
 | 
| 226 | ||
| 227 |   \begin{description}
 | |
| 228 | ||
| 39864 | 229 |   \item Type @{ML_type Proof.context} represents proof contexts.
 | 
| 20449 | 230 | |
| 52788 | 231 |   \item @{ML Proof_Context.init_global}~@{text "thy"} produces a proof
 | 
| 232 |   context derived from @{text "thy"}, initializing all data.
 | |
| 20449 | 233 | |
| 42361 | 234 |   \item @{ML Proof_Context.theory_of}~@{text "ctxt"} selects the
 | 
| 52788 | 235 |   background theory from @{text "ctxt"}.
 | 
| 20449 | 236 | |
| 42361 | 237 |   \item @{ML Proof_Context.transfer}~@{text "thy ctxt"} promotes the
 | 
| 20449 | 238 |   background theory of @{text "ctxt"} to the super theory @{text
 | 
| 239 | "thy"}. | |
| 240 | ||
| 241 |   \end{description}
 | |
| 242 | *} | |
| 243 | ||
| 39832 | 244 | text %mlantiq {*
 | 
| 245 |   \begin{matharray}{rcl}
 | |
| 246 |   @{ML_antiquotation_def "context"} & : & @{text ML_antiquotation} \\
 | |
| 247 |   \end{matharray}
 | |
| 248 | ||
| 249 |   \begin{description}
 | |
| 250 | ||
| 251 |   \item @{text "@{context}"} refers to \emph{the} context at
 | |
| 252 | compile-time --- as abstract value. Independently of (local) theory | |
| 253 | or proof mode, this always produces a meaningful result. | |
| 254 | ||
| 255 | This is probably the most common antiquotation in interactive | |
| 256 | experimentation with ML inside Isar. | |
| 257 | ||
| 258 |   \end{description}
 | |
| 259 | *} | |
| 260 | ||
| 20430 | 261 | |
| 20451 | 262 | subsection {* Generic contexts \label{sec:generic-context} *}
 | 
| 20429 | 263 | |
| 20449 | 264 | text {*
 | 
| 265 | A generic context is the disjoint sum of either a theory or proof | |
| 20451 | 266 | context. Occasionally, this enables uniform treatment of generic | 
| 20450 | 267 | context data, typically extra-logical information. Operations on | 
| 20449 | 268 | generic contexts include the usual injections, partial selections, | 
| 269 | and combinators for lifting operations on either component of the | |
| 270 | disjoint sum. | |
| 271 | ||
| 272 |   Moreover, there are total operations @{text "theory_of"} and @{text
 | |
| 273 | "proof_of"} to convert a generic context into either kind: a theory | |
| 20451 | 274 | can always be selected from the sum, while a proof context might | 
| 34921 | 275 |   have to be constructed by an ad-hoc @{text "init"} operation, which
 | 
| 276 | incurs a small runtime overhead. | |
| 20449 | 277 | *} | 
| 20430 | 278 | |
| 20449 | 279 | text %mlref {*
 | 
| 280 |   \begin{mldecls}
 | |
| 281 |   @{index_ML_type Context.generic} \\
 | |
| 282 |   @{index_ML Context.theory_of: "Context.generic -> theory"} \\
 | |
| 283 |   @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
 | |
| 284 |   \end{mldecls}
 | |
| 285 | ||
| 286 |   \begin{description}
 | |
| 20430 | 287 | |
| 39864 | 288 |   \item Type @{ML_type Context.generic} is the direct sum of @{ML_type
 | 
| 20451 | 289 |   "theory"} and @{ML_type "Proof.context"}, with the datatype
 | 
| 290 |   constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
 | |
| 20449 | 291 | |
| 292 |   \item @{ML Context.theory_of}~@{text "context"} always produces a
 | |
| 293 |   theory from the generic @{text "context"}, using @{ML
 | |
| 42361 | 294 | "Proof_Context.theory_of"} as required. | 
| 20449 | 295 | |
| 296 |   \item @{ML Context.proof_of}~@{text "context"} always produces a
 | |
| 297 |   proof context from the generic @{text "context"}, using @{ML
 | |
| 42361 | 298 | "Proof_Context.init_global"} as required (note that this re-initializes the | 
| 20451 | 299 | context data with each invocation). | 
| 20449 | 300 | |
| 301 |   \end{description}
 | |
| 302 | *} | |
| 20437 | 303 | |
| 20476 | 304 | |
| 305 | subsection {* Context data \label{sec:context-data} *}
 | |
| 20447 | 306 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 307 | 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 | 308 | arbitrary (pure) data. New data types can be declared incrementally | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 309 | at compile time. There are separate declaration mechanisms for any | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 310 | of the three kinds of contexts: theory, proof, generic. | 
| 20449 | 311 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 312 |   \paragraph{Theory data} declarations need to implement the following
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 313 | SML signature: | 
| 20449 | 314 | |
| 315 | \medskip | |
| 316 |   \begin{tabular}{ll}
 | |
| 22869 | 317 |   @{text "\<type> T"} & representing type \\
 | 
| 318 |   @{text "\<val> empty: T"} & empty default value \\
 | |
| 319 |   @{text "\<val> extend: T \<rightarrow> T"} & re-initialize on import \\
 | |
| 320 |   @{text "\<val> merge: T \<times> T \<rightarrow> T"} & join on import \\
 | |
| 20449 | 321 |   \end{tabular}
 | 
| 322 | \medskip | |
| 323 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 324 |   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 | 325 |   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 | 326 |   is acts like a unitary version of @{text "merge"}.
 | 
| 20449 | 327 | |
| 34921 | 328 |   Implementing @{text "merge"} can be tricky.  The general idea is
 | 
| 329 |   that @{text "merge (data\<^sub>1, data\<^sub>2)"} inserts those parts of @{text
 | |
| 330 |   "data\<^sub>2"} into @{text "data\<^sub>1"} that are not yet present, while
 | |
| 331 |   keeping the general order of things.  The @{ML Library.merge}
 | |
| 332 | function on plain lists may serve as canonical template. | |
| 333 | ||
| 334 | Particularly note that shared parts of the data must not be | |
| 335 | duplicated by naive concatenation, or a theory graph that is like a | |
| 336 | chain of diamonds would cause an exponential blowup! | |
| 337 | ||
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 338 |   \paragraph{Proof context data} declarations need to implement the
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 339 | following SML signature: | 
| 20449 | 340 | |
| 341 | \medskip | |
| 342 |   \begin{tabular}{ll}
 | |
| 22869 | 343 |   @{text "\<type> T"} & representing type \\
 | 
| 344 |   @{text "\<val> init: theory \<rightarrow> T"} & produce initial value \\
 | |
| 20449 | 345 |   \end{tabular}
 | 
| 346 | \medskip | |
| 347 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 348 |   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 | 349 | from the given background theory and should be somehow | 
| 34921 | 350 | ``immediate''. Whenever a proof context is initialized, which | 
| 351 |   happens frequently, the the system invokes the @{text "init"}
 | |
| 39821 | 352 |   operation of \emph{all} theory data slots ever declared.  This also
 | 
| 353 | means that one needs to be economic about the total number of proof | |
| 354 | data declarations in the system, i.e.\ each ML module should declare | |
| 355 | at most one, sometimes two data slots for its internal use. | |
| 356 | Repeated data declarations to simulate a record type should be | |
| 357 | avoided! | |
| 20449 | 358 | |
| 20451 | 359 |   \paragraph{Generic data} provides a hybrid interface for both theory
 | 
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 360 |   and proof data.  The @{text "init"} operation for proof contexts is
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 361 | predefined to select the current data value from the background | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 362 | theory. | 
| 20449 | 363 | |
| 39821 | 364 |   \bigskip Any of the above data declarations over type @{text "T"}
 | 
| 365 | result in an ML structure with the following signature: | |
| 20449 | 366 | |
| 367 | \medskip | |
| 368 |   \begin{tabular}{ll}
 | |
| 369 |   @{text "get: context \<rightarrow> T"} \\
 | |
| 370 |   @{text "put: T \<rightarrow> context \<rightarrow> context"} \\
 | |
| 371 |   @{text "map: (T \<rightarrow> T) \<rightarrow> context \<rightarrow> context"} \\
 | |
| 372 |   \end{tabular}
 | |
| 373 | \medskip | |
| 374 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 375 | 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 | 376 | 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 | 377 | 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 | 378 | 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 | 379 | 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 | 380 | abstract values authentically. *} | 
| 20447 | 381 | |
| 20450 | 382 | text %mlref {*
 | 
| 383 |   \begin{mldecls}
 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 384 |   @{index_ML_functor Theory_Data} \\
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 385 |   @{index_ML_functor Proof_Data} \\
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 386 |   @{index_ML_functor Generic_Data} \\
 | 
| 20450 | 387 |   \end{mldecls}
 | 
| 388 | ||
| 389 |   \begin{description}
 | |
| 390 | ||
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 391 |   \item @{ML_functor Theory_Data}@{text "(spec)"} declares data for
 | 
| 20450 | 392 |   type @{ML_type theory} according to the specification provided as
 | 
| 20451 | 393 | argument structure. The resulting structure provides data init and | 
| 394 | access operations as described above. | |
| 20450 | 395 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 396 |   \item @{ML_functor Proof_Data}@{text "(spec)"} is analogous to
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 397 |   @{ML_functor Theory_Data} for type @{ML_type Proof.context}.
 | 
| 20450 | 398 | |
| 33524 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 399 |   \item @{ML_functor Generic_Data}@{text "(spec)"} is analogous to
 | 
| 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 wenzelm parents: 
33174diff
changeset | 400 |   @{ML_functor Theory_Data} for type @{ML_type Context.generic}.
 | 
| 20450 | 401 | |
| 402 |   \end{description}
 | |
| 403 | *} | |
| 404 | ||
| 34928 | 405 | text %mlex {*
 | 
| 406 | The following artificial example demonstrates theory | |
| 407 | data: we maintain a set of terms that are supposed to be wellformed | |
| 408 | wrt.\ the enclosing theory. The public interface is as follows: | |
| 409 | *} | |
| 410 | ||
| 411 | ML {*
 | |
| 412 | signature WELLFORMED_TERMS = | |
| 413 | sig | |
| 414 | val get: theory -> term list | |
| 415 | val add: term -> theory -> theory | |
| 416 | end; | |
| 417 | *} | |
| 418 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 419 | 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 | 420 | 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 | 421 | wrt.\ the given theory. *} | 
| 34928 | 422 | |
| 423 | ML {*
 | |
| 424 | structure Wellformed_Terms: WELLFORMED_TERMS = | |
| 425 | struct | |
| 426 | ||
| 427 | structure Terms = Theory_Data | |
| 428 | ( | |
| 39687 | 429 | type T = term Ord_List.T; | 
| 34928 | 430 | val empty = []; | 
| 431 | val extend = I; | |
| 432 | fun merge (ts1, ts2) = | |
| 39687 | 433 | 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 | 434 | ); | 
| 34928 | 435 | |
| 436 | val get = Terms.get; | |
| 437 | ||
| 438 | fun add raw_t thy = | |
| 39821 | 439 | let | 
| 440 | val t = Sign.cert_term thy raw_t; | |
| 441 | in | |
| 442 | Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy | |
| 443 | end; | |
| 34928 | 444 | |
| 445 | end; | |
| 446 | *} | |
| 447 | ||
| 39864 | 448 | text {* Type @{ML_type "term Ord_List.T"} is used for reasonably
 | 
| 449 | efficient representation of a set of terms: all operations are | |
| 450 | linear in the number of stored elements. Here we assume that users | |
| 451 | of this module do not care about the declaration order, since that | |
| 452 | data structure forces its own arrangement of elements. | |
| 34928 | 453 | |
| 40153 | 454 |   Observe how the @{ML_text merge} operation joins the data slots of
 | 
| 39687 | 455 |   the two constituents: @{ML Ord_List.union} prevents duplication of
 | 
| 34928 | 456 | common data from different branches, thus avoiding the danger of | 
| 39821 | 457 | exponential blowup. Plain list append etc.\ must never be used for | 
| 458 | theory data merges! | |
| 34928 | 459 | |
| 460 | \medskip Our intended invariant is achieved as follows: | |
| 461 |   \begin{enumerate}
 | |
| 462 | ||
| 463 |   \item @{ML Wellformed_Terms.add} only admits terms that have passed
 | |
| 464 |   the @{ML Sign.cert_term} check of the given theory at that point.
 | |
| 465 | ||
| 466 |   \item Wellformedness in the sense of @{ML Sign.cert_term} is
 | |
| 467 | monotonic wrt.\ the sub-theory relation. So our data can move | |
| 468 | upwards in the hierarchy (via extension or merges), and maintain | |
| 469 | wellformedness without further checks. | |
| 470 | ||
| 471 |   \end{enumerate}
 | |
| 472 | ||
| 473 | Note that all basic operations of the inference kernel (which | |
| 474 |   includes @{ML Sign.cert_term}) observe this monotonicity principle,
 | |
| 475 | but other user-space tools don't. For example, fully-featured | |
| 476 |   type-inference via @{ML Syntax.check_term} (cf.\
 | |
| 477 |   \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
 | |
| 478 | background theory, since constraints of term constants can be | |
| 39821 | 479 | modified by later declarations, for example. | 
| 34928 | 480 | |
| 481 | In most cases, user-space context data does not have to take such | |
| 482 | invariants too seriously. The situation is different in the | |
| 483 | implementation of the inference kernel itself, which uses the very | |
| 484 | same data mechanisms for types, constants, axioms etc. | |
| 485 | *} | |
| 486 | ||
| 20447 | 487 | |
| 39865 | 488 | subsection {* Configuration options \label{sec:config-options} *}
 | 
| 489 | ||
| 490 | text {* A \emph{configuration option} is a named optional value of
 | |
| 491 | some basic type (Boolean, integer, string) that is stored in the | |
| 492 | context. It is a simple application of general context data | |
| 493 |   (\secref{sec:context-data}) that is sufficiently common to justify
 | |
| 494 | customized setup, which includes some concrete declarations for | |
| 495 | end-users using existing notation for attributes (cf.\ | |
| 496 |   \secref{sec:attributes}).
 | |
| 497 | ||
| 498 |   For example, the predefined configuration option @{attribute
 | |
| 499 | show_types} controls output of explicit type constraints for | |
| 39876 
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
 wenzelm parents: 
39866diff
changeset | 500 |   variables in printed terms (cf.\ \secref{sec:read-print}).  Its
 | 
| 39865 | 501 | value can be modified within Isar text like this: | 
| 502 | *} | |
| 503 | ||
| 504 | declare [[show_types = false]] | |
| 505 |   -- {* declaration within (local) theory context *}
 | |
| 506 | ||
| 40964 | 507 | notepad | 
| 508 | begin | |
| 39865 | 509 | note [[show_types = true]] | 
| 510 |     -- {* declaration within proof (forward mode) *}
 | |
| 511 | term x | |
| 512 | ||
| 513 | have "x = x" | |
| 514 | using [[show_types = false]] | |
| 515 |       -- {* declaration within proof (backward mode) *}
 | |
| 516 | .. | |
| 40964 | 517 | end | 
| 39865 | 518 | |
| 519 | text {* Configuration options that are not set explicitly hold a
 | |
| 520 | default value that can depend on the application context. This | |
| 521 | allows to retrieve the value from another slot within the context, | |
| 522 | or fall back on a global preference mechanism, for example. | |
| 523 | ||
| 524 | The operations to declare configuration options and get/map their | |
| 525 | values are modeled as direct replacements for historic global | |
| 526 | references, only that the context is made explicit. This allows | |
| 527 | easy configuration of tools, without relying on the execution order | |
| 528 | as required for old-style mutable references. *} | |
| 529 | ||
| 530 | text %mlref {*
 | |
| 531 |   \begin{mldecls}
 | |
| 532 |   @{index_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\
 | |
| 533 |   @{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 | 534 |   @{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 | 535 | bool Config.T"} \\ | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 536 |   @{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 | 537 | int Config.T"} \\ | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 538 |   @{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 | 539 | real Config.T"} \\ | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 540 |   @{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 | 541 | string Config.T"} \\ | 
| 39865 | 542 |   \end{mldecls}
 | 
| 543 | ||
| 544 |   \begin{description}
 | |
| 545 | ||
| 546 |   \item @{ML Config.get}~@{text "ctxt config"} gets the value of
 | |
| 547 |   @{text "config"} in the given context.
 | |
| 548 | ||
| 549 |   \item @{ML Config.map}~@{text "config f ctxt"} updates the context
 | |
| 550 |   by updating the value of @{text "config"}.
 | |
| 551 | ||
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 552 |   \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 | 553 |   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 | 554 |   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 | 555 |   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 | 556 | 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 | 557 | 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 | 558 | concrete syntax. | 
| 39865 | 559 | |
| 40291 | 560 |   \item @{ML Attrib.config_int}, @{ML Attrib.config_real}, and @{ML
 | 
| 561 |   Attrib.config_string} work like @{ML Attrib.config_bool}, but for
 | |
| 562 |   types @{ML_type int} and @{ML_type string}, respectively.
 | |
| 39865 | 563 | |
| 564 |   \end{description}
 | |
| 565 | *} | |
| 566 | ||
| 567 | text %mlex {* The following example shows how to declare and use a
 | |
| 568 |   Boolean configuration option called @{text "my_flag"} with constant
 | |
| 569 |   default value @{ML false}.  *}
 | |
| 570 | ||
| 571 | ML {*
 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 572 | val my_flag = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 573 |     Attrib.setup_config_bool @{binding my_flag} (K false)
 | 
| 39865 | 574 | *} | 
| 575 | ||
| 576 | text {* Now the user can refer to @{attribute my_flag} in
 | |
| 40126 | 577 | declarations, while ML tools can retrieve the current value from the | 
| 39865 | 578 |   context via @{ML Config.get}.  *}
 | 
| 579 | ||
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 580 | ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
 | 
| 39865 | 581 | |
| 582 | declare [[my_flag = true]] | |
| 583 | ||
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 584 | ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
 | 
| 39865 | 585 | |
| 40964 | 586 | notepad | 
| 587 | begin | |
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 588 |   {
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 589 | note [[my_flag = false]] | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 590 |     ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 591 | } | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 592 |   ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
 | 
| 40964 | 593 | end | 
| 39865 | 594 | |
| 40291 | 595 | text {* Here is another example involving ML type @{ML_type real}
 | 
| 596 | (floating-point numbers). *} | |
| 597 | ||
| 598 | ML {*
 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 599 | val airspeed_velocity = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42510diff
changeset | 600 |     Attrib.setup_config_real @{binding airspeed_velocity} (K 0.0)
 | 
| 40291 | 601 | *} | 
| 602 | ||
| 40296 | 603 | declare [[airspeed_velocity = 10]] | 
| 40291 | 604 | declare [[airspeed_velocity = 9.9]] | 
| 605 | ||
| 39865 | 606 | |
| 26872 | 607 | section {* Names \label{sec:names} *}
 | 
| 20451 | 608 | |
| 34925 | 609 | text {* In principle, a name is just a string, but there are various
 | 
| 610 | conventions for representing additional structure. For example, | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 611 |   ``@{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 | 612 |   qualifier @{text "Foo.bar"} and base name @{text "baz"}.  The
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 613 | individual constituents of a name may have further substructure, | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 614 | e.g.\ the string ``\verb,\,\verb,<alpha>,'' encodes as a single | 
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
51686diff
changeset | 615 |   symbol (\secref{sec:symbols}).
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 616 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 617 | \medskip Subsequently, we shall introduce specific categories of | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 618 | names. Roughly speaking these correspond to logical entities as | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 619 | follows: | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 620 |   \begin{itemize}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 621 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 622 |   \item Basic names (\secref{sec:basic-name}): free and bound
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 623 | variables. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 624 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 625 |   \item Indexed names (\secref{sec:indexname}): schematic variables.
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 626 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 627 |   \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 | 628 | (type constructors, term constants, other concepts defined in user | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 629 | space). Such entities are typically managed via name spaces | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 630 |   (\secref{sec:name-space}).
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 631 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 632 |   \end{itemize}
 | 
| 20451 | 633 | *} | 
| 20437 | 634 | |
| 635 | ||
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 636 | subsection {* Basic names \label{sec:basic-name} *}
 | 
| 20476 | 637 | |
| 638 | text {*
 | |
| 639 |   A \emph{basic name} essentially consists of a single Isabelle
 | |
| 640 | identifier. There are conventions to mark separate classes of basic | |
| 29761 | 641 | names, by attaching a suffix of underscores: one underscore means | 
| 642 |   \emph{internal name}, two underscores means \emph{Skolem name},
 | |
| 643 |   three underscores means \emph{internal Skolem name}.
 | |
| 20476 | 644 | |
| 645 |   For example, the basic name @{text "foo"} has the internal version
 | |
| 646 |   @{text "foo_"}, with Skolem versions @{text "foo__"} and @{text
 | |
| 647 | "foo___"}, respectively. | |
| 648 | ||
| 20488 | 649 | These special versions provide copies of the basic name space, apart | 
| 650 | from anything that normally appears in the user text. For example, | |
| 651 | system generated variables in Isar proof contexts are usually marked | |
| 34926 | 652 |   as internal, which prevents mysterious names like @{text "xaa"} to
 | 
| 653 | appear in human-readable text. | |
| 20476 | 654 | |
| 20488 | 655 | \medskip Manipulating binding scopes often requires on-the-fly | 
| 656 |   renamings.  A \emph{name context} contains a collection of already
 | |
| 657 |   used names.  The @{text "declare"} operation adds names to the
 | |
| 658 | context. | |
| 20476 | 659 | |
| 20488 | 660 |   The @{text "invents"} operation derives a number of fresh names from
 | 
| 661 | a given starting point. For example, the first three names derived | |
| 662 |   from @{text "a"} are @{text "a"}, @{text "b"}, @{text "c"}.
 | |
| 20476 | 663 | |
| 664 |   The @{text "variants"} operation produces fresh names by
 | |
| 20488 | 665 |   incrementing tentative names as base-26 numbers (with digits @{text
 | 
| 666 |   "a..z"}) until all clashes are resolved.  For example, name @{text
 | |
| 667 |   "foo"} results in variants @{text "fooa"}, @{text "foob"}, @{text
 | |
| 668 |   "fooc"}, \dots, @{text "fooaa"}, @{text "fooab"} etc.; each renaming
 | |
| 669 | step picks the next unused variant from this sequence. | |
| 20476 | 670 | *} | 
| 671 | ||
| 672 | text %mlref {*
 | |
| 673 |   \begin{mldecls}
 | |
| 674 |   @{index_ML Name.internal: "string -> string"} \\
 | |
| 20547 | 675 |   @{index_ML Name.skolem: "string -> string"} \\
 | 
| 676 |   \end{mldecls}
 | |
| 677 |   \begin{mldecls}
 | |
| 20476 | 678 |   @{index_ML_type Name.context} \\
 | 
| 679 |   @{index_ML Name.context: Name.context} \\
 | |
| 680 |   @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 681 |   @{index_ML Name.invent: "Name.context -> string -> int -> string list"} \\
 | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42616diff
changeset | 682 |   @{index_ML Name.variant: "string -> Name.context -> string * Name.context"} \\
 | 
| 20476 | 683 |   \end{mldecls}
 | 
| 34926 | 684 |   \begin{mldecls}
 | 
| 685 |   @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
 | |
| 686 |   \end{mldecls}
 | |
| 20476 | 687 | |
| 688 |   \begin{description}
 | |
| 689 | ||
| 690 |   \item @{ML Name.internal}~@{text "name"} produces an internal name
 | |
| 691 | by adding one underscore. | |
| 692 | ||
| 693 |   \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by
 | |
| 694 | adding two underscores. | |
| 695 | ||
| 39864 | 696 |   \item Type @{ML_type Name.context} represents the context of already
 | 
| 697 |   used names; the initial value is @{ML "Name.context"}.
 | |
| 20476 | 698 | |
| 20488 | 699 |   \item @{ML Name.declare}~@{text "name"} enters a used name into the
 | 
| 700 | context. | |
| 20437 | 701 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 702 |   \item @{ML Name.invent}~@{text "context name n"} produces @{text
 | 
| 20488 | 703 |   "n"} fresh names derived from @{text "name"}.
 | 
| 704 | ||
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42616diff
changeset | 705 |   \item @{ML Name.variant}~@{text "name context"} produces a fresh
 | 
| 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42616diff
changeset | 706 |   variant of @{text "name"}; the result is declared to the context.
 | 
| 20476 | 707 | |
| 34926 | 708 |   \item @{ML Variable.names_of}~@{text "ctxt"} retrieves the context
 | 
| 709 | of declared type and term variable names. Projecting a proof | |
| 710 | context down to a primitive name context is occasionally useful when | |
| 711 | invoking lower-level operations. Regular management of ``fresh | |
| 712 |   variables'' is done by suitable operations of structure @{ML_struct
 | |
| 713 | Variable}, which is also able to provide an official status of | |
| 714 | ``locally fixed variable'' within the logical environment (cf.\ | |
| 715 |   \secref{sec:variables}).
 | |
| 716 | ||
| 20476 | 717 |   \end{description}
 | 
| 718 | *} | |
| 719 | ||
| 39857 | 720 | text %mlex {* The following simple examples demonstrate how to produce
 | 
| 721 |   fresh names from the initial @{ML Name.context}. *}
 | |
| 722 | ||
| 723 | ML {*
 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 724 | val list1 = Name.invent Name.context "a" 5; | 
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 725 |   @{assert} (list1 = ["a", "b", "c", "d", "e"]);
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 726 | |
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 727 | val list2 = | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42616diff
changeset | 728 | #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] Name.context); | 
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 729 |   @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]);
 | 
| 39857 | 730 | *} | 
| 731 | ||
| 40126 | 732 | 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 | 733 | follows. *} | 
| 39857 | 734 | |
| 735 | locale ex = fixes a b c :: 'a | |
| 736 | begin | |
| 737 | ||
| 738 | ML {*
 | |
| 739 |   val names = Variable.names_of @{context};
 | |
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 740 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 741 | val list1 = Name.invent names "a" 5; | 
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 742 |   @{assert} (list1 = ["d", "e", "f", "g", "h"]);
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 743 | |
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 744 | val list2 = | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42616diff
changeset | 745 | #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] names); | 
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39865diff
changeset | 746 |   @{assert} (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]);
 | 
| 39857 | 747 | *} | 
| 748 | ||
| 749 | end | |
| 750 | ||
| 20476 | 751 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 752 | subsection {* Indexed names \label{sec:indexname} *}
 | 
| 20476 | 753 | |
| 754 | text {*
 | |
| 755 |   An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic
 | |
| 20488 | 756 | name and a natural number. This representation allows efficient | 
| 757 | renaming by incrementing the second component only. The canonical | |
| 758 | way to rename two collections of indexnames apart from each other is | |
| 759 |   this: determine the maximum index @{text "maxidx"} of the first
 | |
| 760 | collection, then increment all indexes of the second collection by | |
| 761 |   @{text "maxidx + 1"}; the maximum index of an empty collection is
 | |
| 762 |   @{text "-1"}.
 | |
| 20476 | 763 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 764 | 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 | 765 |   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 | 766 |   name @{text "x"}.
 | 
| 20488 | 767 | |
| 768 | \medskip Isabelle syntax observes the following rules for | |
| 769 |   representing an indexname @{text "(x, i)"} as a packed string:
 | |
| 20476 | 770 | |
| 771 |   \begin{itemize}
 | |
| 772 | ||
| 20479 | 773 |   \item @{text "?x"} if @{text "x"} does not end with a digit and @{text "i = 0"},
 | 
| 20476 | 774 | |
| 775 |   \item @{text "?xi"} if @{text "x"} does not end with a digit,
 | |
| 776 | ||
| 20488 | 777 |   \item @{text "?x.i"} otherwise.
 | 
| 20476 | 778 | |
| 779 |   \end{itemize}
 | |
| 20470 | 780 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 781 | Indexnames may acquire large index numbers after several maxidx | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 782 | shifts have been applied. Results are usually normalized towards | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 783 |   @{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 | 784 | This works by producing variants of the corresponding basic name | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 785 |   components.  For example, the collection @{text "?x1, ?x7, ?x42"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 786 |   becomes @{text "?x, ?xa, ?xb"}.
 | 
| 20476 | 787 | *} | 
| 788 | ||
| 789 | text %mlref {*
 | |
| 790 |   \begin{mldecls}
 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 791 |   @{index_ML_type indexname: "string * int"} \\
 | 
| 20476 | 792 |   \end{mldecls}
 | 
| 793 | ||
| 794 |   \begin{description}
 | |
| 795 | ||
| 39864 | 796 |   \item Type @{ML_type indexname} represents indexed names.  This is
 | 
| 797 |   an abbreviation for @{ML_type "string * int"}.  The second component
 | |
| 798 |   is usually non-negative, except for situations where @{text "(x,
 | |
| 799 | -1)"} is used to inject basic names into this type. Other negative | |
| 34926 | 800 | indexes should not be used. | 
| 20476 | 801 | |
| 802 |   \end{description}
 | |
| 803 | *} | |
| 804 | ||
| 805 | ||
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 806 | subsection {* Long names \label{sec:long-name} *}
 | 
| 20476 | 807 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 808 | 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 | 809 | components. The packed representation uses a dot as separator, as | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 810 |   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 | 811 |   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 | 812 | 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 | 813 | named entity while passing through some nested block-structure, | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 814 | 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 | 815 | discipline. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 816 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 817 |   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 | 818 |   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 | 819 |   within a global structure @{text "A"}.  In practice, long names
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 820 | 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 | 821 | not make any assumptions about the particular structure of long | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 822 | names! | 
| 20437 | 823 | |
| 20476 | 824 | 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 | 825 | entities, or entities that are not entered into the corresponding | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 826 | 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 | 827 | long names map empty names again to empty names. | 
| 20437 | 828 | *} | 
| 829 | ||
| 20476 | 830 | text %mlref {*
 | 
| 831 |   \begin{mldecls}
 | |
| 30365 | 832 |   @{index_ML Long_Name.base_name: "string -> string"} \\
 | 
| 833 |   @{index_ML Long_Name.qualifier: "string -> string"} \\
 | |
| 834 |   @{index_ML Long_Name.append: "string -> string -> string"} \\
 | |
| 835 |   @{index_ML Long_Name.implode: "string list -> string"} \\
 | |
| 836 |   @{index_ML Long_Name.explode: "string -> string list"} \\
 | |
| 20547 | 837 |   \end{mldecls}
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 838 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 839 |   \begin{description}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 840 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 841 |   \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 | 842 | of a long name. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 843 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 844 |   \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 845 | of a long name. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 846 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52788diff
changeset | 847 |   \item @{ML Long_Name.append}~@{text "name\<^sub>1 name\<^sub>2"} appends two long
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 848 | names. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 849 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 850 |   \item @{ML Long_Name.implode}~@{text "names"} and @{ML
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 851 |   Long_Name.explode}~@{text "name"} convert between the packed string
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 852 | representation and the explicit list form of long names. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 853 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 854 |   \end{description}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 855 | *} | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 856 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 857 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 858 | subsection {* Name spaces \label{sec:name-space} *}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 859 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 860 | 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 | 861 | together with a mapping between partially qualified external names | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 862 | and fully qualified internal names (in both directions). Note that | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 863 |   the corresponding @{text "intern"} and @{text "extern"} operations
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 864 |   are mostly used for parsing and printing only!  The @{text
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 865 | "declare"} operation augments a name space according to the accesses | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 866 | 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 | 867 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 868 |   \medskip A @{text "binding"} specifies details about the prospective
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 869 | 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 | 870 | base name, prefixes for qualification (separate ones for system | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 871 | infrastructure and user-space mechanisms), a slot for the original | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 872 | source position, and some additional flags. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 873 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 874 |   \medskip A @{text "naming"} provides some additional details for
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 875 | 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 | 876 |   implicit in the theory or proof context.  The @{text "full"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 877 | operation (and its variants for different context types) produces a | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 878 | 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 | 879 | main equation of this ``chemical reaction'' when binding new | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 880 | entities in a context is as follows: | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 881 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 882 | \medskip | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 883 |   \begin{tabular}{l}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 884 |   @{text "binding + naming \<longrightarrow> long name + name space accesses"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 885 |   \end{tabular}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 886 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 887 | \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 | 888 | 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 | 889 | 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 | 890 | 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 | 891 |   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 | 892 | uniformly for a constant, type constructor, and type class. | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 893 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 894 | There are common schemes to name derived entities systematically | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 895 | 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 | 896 |   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 | 897 |   constant @{text "c"}.  This technique of mapping names from one
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 898 | 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 | 899 | In particular, theorem names derived from a type constructor or type | 
| 39839 | 900 | class should get an additional suffix in addition to the usual | 
| 901 | qualification. This leads to the following conventions for derived | |
| 902 | names: | |
| 903 | ||
| 904 | \medskip | |
| 905 |   \begin{tabular}{ll}
 | |
| 906 | logical entity & fact name \\\hline | |
| 907 |   constant @{text "c"} & @{text "c.intro"} \\
 | |
| 908 |   type @{text "c"} & @{text "c_type.intro"} \\
 | |
| 909 |   class @{text "c"} & @{text "c_class.intro"} \\
 | |
| 910 |   \end{tabular}
 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 911 | *} | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 912 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 913 | text %mlref {*
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 914 |   \begin{mldecls}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 915 |   @{index_ML_type binding} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 916 |   @{index_ML Binding.empty: binding} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 917 |   @{index_ML Binding.name: "string -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 918 |   @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 919 |   @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 920 |   @{index_ML Binding.conceal: "binding -> binding"} \\
 | 
| 43547 
f3a8476285c6
clarified Binding.pretty/print: no quotes, only markup -- Binding.str_of is rendered obsolete;
 wenzelm parents: 
43329diff
changeset | 921 |   @{index_ML Binding.print: "binding -> string"} \\
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 922 |   \end{mldecls}
 | 
| 20547 | 923 |   \begin{mldecls}
 | 
| 33174 | 924 |   @{index_ML_type Name_Space.naming} \\
 | 
| 925 |   @{index_ML Name_Space.default_naming: Name_Space.naming} \\
 | |
| 926 |   @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
 | |
| 927 |   @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
 | |
| 20547 | 928 |   \end{mldecls}
 | 
| 929 |   \begin{mldecls}
 | |
| 33174 | 930 |   @{index_ML_type Name_Space.T} \\
 | 
| 931 |   @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
 | |
| 932 |   @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
 | |
| 47174 | 933 |   @{index_ML Name_Space.declare: "Context.generic -> bool ->
 | 
| 934 | binding -> Name_Space.T -> string * Name_Space.T"} \\ | |
| 33174 | 935 |   @{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 | 936 |   @{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 | 937 |   @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
 | 
| 20476 | 938 |   \end{mldecls}
 | 
| 20437 | 939 | |
| 20476 | 940 |   \begin{description}
 | 
| 941 | ||
| 39864 | 942 |   \item Type @{ML_type binding} represents the abstract concept of
 | 
| 943 | name bindings. | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 944 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 945 |   \item @{ML Binding.empty} is the empty binding.
 | 
| 20476 | 946 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 947 |   \item @{ML Binding.name}~@{text "name"} produces a binding with base
 | 
| 39832 | 948 |   name @{text "name"}.  Note that this lacks proper source position
 | 
| 949 |   information; see also the ML antiquotation @{ML_antiquotation
 | |
| 950 | binding}. | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 951 | |
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 952 |   \item @{ML Binding.qualify}~@{text "mandatory name binding"}
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 953 |   prefixes qualifier @{text "name"} to @{text "binding"}.  The @{text
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 954 | "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 | 955 |   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 | 956 | 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 | 957 | derived specification mechanisms. | 
| 20437 | 958 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 959 |   \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 | 960 | affects the system prefix. This part of extra qualification is | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 961 | typically used in the infrastructure for modular specifications, | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 962 |   notably ``local theory targets'' (see also \chref{ch:local-theory}).
 | 
| 20437 | 963 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 964 |   \item @{ML Binding.conceal}~@{text "binding"} indicates that the
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 965 | binding shall refer to an entity that serves foundational purposes | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 966 | only. This flag helps to mark implementation details of | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 967 | specification mechanism etc. Other tools should not depend on the | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 968 |   particulars of concealed entities (cf.\ @{ML
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 969 | Name_Space.is_concealed}). | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 970 | |
| 43547 
f3a8476285c6
clarified Binding.pretty/print: no quotes, only markup -- Binding.str_of is rendered obsolete;
 wenzelm parents: 
43329diff
changeset | 971 |   \item @{ML Binding.print}~@{text "binding"} produces a string
 | 
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 972 | representation for human-readable output, together with some formal | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 973 | markup that might get used in GUI front-ends, for example. | 
| 20476 | 974 | |
| 39864 | 975 |   \item Type @{ML_type Name_Space.naming} represents the abstract
 | 
| 976 | concept of a naming policy. | |
| 20437 | 977 | |
| 33174 | 978 |   \item @{ML Name_Space.default_naming} is the default naming policy.
 | 
| 20476 | 979 | In a theory context, this is usually augmented by a path prefix | 
| 980 | consisting of the theory name. | |
| 981 | ||
| 33174 | 982 |   \item @{ML Name_Space.add_path}~@{text "path naming"} augments the
 | 
| 20488 | 983 | naming policy by extending its path component. | 
| 20437 | 984 | |
| 33174 | 985 |   \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 | 986 | name binding (usually a basic name) into the fully qualified | 
| 29008 | 987 | internal name, according to the given naming policy. | 
| 20476 | 988 | |
| 39864 | 989 |   \item Type @{ML_type Name_Space.T} represents name spaces.
 | 
| 20476 | 990 | |
| 33174 | 991 |   \item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52788diff
changeset | 992 | "(space\<^sub>1, space\<^sub>2)"} are the canonical operations for | 
| 20488 | 993 | maintaining name spaces according to theory data management | 
| 33174 | 994 |   (\secref{sec:context-data}); @{text "kind"} is a formal comment
 | 
| 995 | to characterize the purpose of a name space. | |
| 20437 | 996 | |
| 47174 | 997 |   \item @{ML Name_Space.declare}~@{text "context strict binding
 | 
| 33174 | 998 | space"} enters a name binding as fully qualified internal name into | 
| 47174 | 999 | the name space, using the naming of the context. | 
| 20476 | 1000 | |
| 33174 | 1001 |   \item @{ML Name_Space.intern}~@{text "space name"} internalizes a
 | 
| 20476 | 1002 | (partially qualified) external name. | 
| 20437 | 1003 | |
| 20488 | 1004 | This operation is mostly for parsing! Note that fully qualified | 
| 20476 | 1005 |   names stemming from declarations are produced via @{ML
 | 
| 33174 | 1006 |   "Name_Space.full_name"} and @{ML "Name_Space.declare"}
 | 
| 29008 | 1007 |   (or their derivatives for @{ML_type theory} and
 | 
| 20488 | 1008 |   @{ML_type Proof.context}).
 | 
| 20437 | 1009 | |
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
40964diff
changeset | 1010 |   \item @{ML Name_Space.extern}~@{text "ctxt space name"} externalizes a
 | 
| 20476 | 1011 | (fully qualified) internal name. | 
| 1012 | ||
| 30281 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 wenzelm parents: 
30272diff
changeset | 1013 | 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 | 1014 | the precise result too much. | 
| 20476 | 1015 | |
| 34927 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1016 |   \item @{ML Name_Space.is_concealed}~@{text "space name"} indicates
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1017 |   whether @{text "name"} refers to a strictly private entity that
 | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1018 | other tools are supposed to ignore! | 
| 
c4c02ac736a6
more details on long names, binding/naming, name space;
 wenzelm parents: 
34926diff
changeset | 1019 | |
| 20476 | 1020 |   \end{description}
 | 
| 1021 | *} | |
| 30272 | 1022 | |
| 39832 | 1023 | text %mlantiq {*
 | 
| 1024 |   \begin{matharray}{rcl}
 | |
| 1025 |   @{ML_antiquotation_def "binding"} & : & @{text ML_antiquotation} \\
 | |
| 1026 |   \end{matharray}
 | |
| 1027 | ||
| 42510 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 1028 |   @{rail "
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 1029 |   @@{ML_antiquotation binding} name
 | 
| 
b9c106763325
use @{rail} antiquotation (with some nested markup);
 wenzelm parents: 
42401diff
changeset | 1030 | "} | 
| 39832 | 1031 | |
| 1032 |   \begin{description}
 | |
| 1033 | ||
| 1034 |   \item @{text "@{binding name}"} produces a binding with base name
 | |
| 1035 |   @{text "name"} and the source position taken from the concrete
 | |
| 1036 | syntax of this antiquotation. In many situations this is more | |
| 1037 |   appropriate than the more basic @{ML Binding.name} function.
 | |
| 1038 | ||
| 1039 |   \end{description}
 | |
| 1040 | *} | |
| 1041 | ||
| 39833 | 1042 | text %mlex {* The following example yields the source position of some
 | 
| 40126 | 1043 | concrete binding inlined into the text: | 
| 39833 | 1044 | *} | 
| 1045 | ||
| 1046 | ML {* Binding.pos_of @{binding here} *}
 | |
| 1047 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1048 | text {* \medskip That position can be also printed in a message as
 | 
| 40126 | 1049 | follows: *} | 
| 39833 | 1050 | |
| 1051 | ML_command {*
 | |
| 1052 | writeln | |
| 48992 | 1053 |     ("Look here" ^ Position.here (Binding.pos_of @{binding here}))
 | 
| 39833 | 1054 | *} | 
| 1055 | ||
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39857diff
changeset | 1056 | 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 | 1057 | opposed to raw specifications of base names: the system can use this | 
| 40126 | 1058 | additional information for feedback given to the user (error | 
| 1059 | messages etc.). *} | |
| 39833 | 1060 | |
| 18537 | 1061 | end |