| author | wenzelm | 
| Sat, 20 May 2023 22:41:37 +0200 | |
| changeset 78090 | 79ad3181071b | 
| parent 76987 | 4c275405faae | 
| child 81116 | 0fb1e2dd4122 | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 2 | ||
| 29755 | 3 | theory Local_Theory | 
| 4 | imports Base | |
| 5 | begin | |
| 18537 | 6 | |
| 58618 | 7 | chapter \<open>Local theory specifications \label{ch:local-theory}\<close>
 | 
| 29759 | 8 | |
| 58618 | 9 | text \<open> | 
| 61854 | 10 | A \<^emph>\<open>local theory\<close> combines aspects of both theory and proof context (cf.\ | 
| 11 |   \secref{sec:context}), such that definitional specifications may be given
 | |
| 12 | relatively to parameters and assumptions. A local theory is represented as a | |
| 13 | regular proof context, augmented by administrative data about the \<^emph>\<open>target | |
| 61477 | 14 | context\<close>. | 
| 29764 | 15 | |
| 61854 | 16 | The target is usually derived from the background theory by adding local | 
| 17 | \<open>\<FIX>\<close> and \<open>\<ASSUME>\<close> elements, plus suitable modifications of | |
| 18 | non-logical context data (e.g.\ a special type-checking discipline). Once | |
| 19 | initialized, the target is ready to absorb definitional primitives: | |
| 20 | \<open>\<DEFINE>\<close> for terms and \<open>\<NOTE>\<close> for theorems. Such definitions may get | |
| 21 | transformed in a target-specific way, but the programming interface hides | |
| 22 | such details. | |
| 29764 | 23 | |
| 24 | Isabelle/Pure provides target mechanisms for locales, type-classes, | |
| 61854 | 25 | type-class instantiations, and general overloading. In principle, users can | 
| 26 | implement new targets as well, but this rather arcane discipline is beyond | |
| 27 | the scope of this manual. In contrast, implementing derived definitional | |
| 28 | packages to be used within a local theory context is quite easy: the | |
| 29 | interfaces are even simpler and more abstract than the underlying primitives | |
| 30 | for raw theories. | |
| 29764 | 31 | |
| 61854 | 32 | Many definitional packages for local theories are available in Isabelle. | 
| 33 | Although a few old packages only work for global theories, the standard way | |
| 34 | of implementing definitional packages in Isabelle is via the local theory | |
| 35 | interface. | |
| 58618 | 36 | \<close> | 
| 29764 | 37 | |
| 38 | ||
| 58618 | 39 | section \<open>Definitional elements\<close> | 
| 18537 | 40 | |
| 58618 | 41 | text \<open> | 
| 61854 | 42 | There are separate elements \<open>\<DEFINE> c \<equiv> t\<close> for terms, and \<open>\<NOTE> b = | 
| 43 | thm\<close> for theorems. Types are treated implicitly, according to Hindley-Milner | |
| 44 |   discipline (cf.\ \secref{sec:variables}). These definitional primitives
 | |
| 45 | essentially act like \<open>let\<close>-bindings within a local context that may already | |
| 46 | contain earlier \<open>let\<close>-bindings and some initial \<open>\<lambda>\<close>-bindings. Thus we gain | |
| 47 | \<^emph>\<open>dependent definitions\<close> that are relative to an initial axiomatic context. | |
| 48 | The following diagram illustrates this idea of axiomatic elements versus | |
| 29765 | 49 | definitional elements: | 
| 29764 | 50 | |
| 51 |   \begin{center}
 | |
| 52 |   \begin{tabular}{|l|l|l|}
 | |
| 53 | \hline | |
| 61493 | 54 | & \<open>\<lambda>\<close>-binding & \<open>let\<close>-binding \\ | 
| 29764 | 55 | \hline | 
| 61493 | 56 | types & fixed \<open>\<alpha>\<close> & arbitrary \<open>\<beta>\<close> \\ | 
| 57 | terms & \<open>\<FIX> x :: \<tau>\<close> & \<open>\<DEFINE> c \<equiv> t\<close> \\ | |
| 58 | theorems & \<open>\<ASSUME> a: A\<close> & \<open>\<NOTE> b = \<^BG>B\<^EN>\<close> \\ | |
| 29764 | 59 | \hline | 
| 60 |   \end{tabular}
 | |
| 61 |   \end{center}
 | |
| 62 | ||
| 61854 | 63 | A user package merely needs to produce suitable \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> | 
| 64 | elements according to the application. For example, a package for inductive | |
| 65 | definitions might first \<open>\<DEFINE>\<close> a certain predicate as some fixed-point | |
| 66 | construction, then \<open>\<NOTE>\<close> a proven result about monotonicity of the | |
| 29765 | 67 | functor involved here, and then produce further derived concepts via | 
| 61493 | 68 | additional \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> elements. | 
| 29764 | 69 | |
| 61854 | 70 | The cumulative sequence of \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> produced at package | 
| 71 | runtime is managed by the local theory infrastructure by means of an | |
| 72 | \<^emph>\<open>auxiliary context\<close>. Thus the system holds up the impression of working | |
| 73 | within a fully abstract situation with hypothetical entities: \<open>\<DEFINE> c \<equiv> | |
| 74 | t\<close> always results in a literal fact \<open>\<^BG>c \<equiv> t\<^EN>\<close>, where \<open>c\<close> is a | |
| 75 | fixed variable \<open>c\<close>. The details about global constants, name spaces etc. are | |
| 76 | handled internally. | |
| 29764 | 77 | |
| 61854 | 78 | So the general structure of a local theory is a sandwich of three layers: | 
| 29764 | 79 | |
| 80 |   \begin{center}
 | |
| 81 |   \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}}
 | |
| 82 |   \end{center}
 | |
| 83 | ||
| 61854 | 84 | When a definitional package is finished, the auxiliary context is reset to | 
| 85 | the target context. The target now holds definitions for terms and theorems | |
| 86 | that stem from the hypothetical \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> elements, | |
| 76987 | 87 | transformed by the particular target policy (see \<^cite>\<open>\<open>\S4--5\<close> in | 
| 88 | "Haftmann-Wenzel:2009"\<close> for details). | |
| 61854 | 89 | \<close> | 
| 29764 | 90 | |
| 58618 | 91 | text %mlref \<open> | 
| 29764 | 92 |   \begin{mldecls}
 | 
| 73764 | 93 |   @{define_ML_type local_theory = Proof.context} \\
 | 
| 94 |   @{define_ML Named_Target.init: "string list -> string -> theory -> local_theory"} \\[1ex]
 | |
| 95 |   @{define_ML Local_Theory.define: "(binding * mixfix) * (Attrib.binding * term) ->
 | |
| 33834 
7c06e19f717c
adapted local theory operations -- eliminated odd kind;
 wenzelm parents: 
33672diff
changeset | 96 | local_theory -> (term * (string * thm)) * local_theory"} \\ | 
| 73764 | 97 |   @{define_ML Local_Theory.note: "Attrib.binding * thm list ->
 | 
| 33672 | 98 | local_theory -> (string * thm list) * local_theory"} \\ | 
| 29764 | 99 |   \end{mldecls}
 | 
| 100 | ||
| 69597 | 101 | \<^descr> Type \<^ML_type>\<open>local_theory\<close> represents local theories. Although this is | 
| 102 | merely an alias for \<^ML_type>\<open>Proof.context\<close>, it is semantically a subtype | |
| 103 | of the same: a \<^ML_type>\<open>local_theory\<close> holds target information as special | |
| 104 | context data. Subtyping means that any value \<open>lthy:\<close>~\<^ML_type>\<open>local_theory\<close> | |
| 105 | can be also used with operations on expecting a regular \<open>ctxt:\<close>~\<^ML_type>\<open>Proof.context\<close>. | |
| 29765 | 106 | |
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
69597diff
changeset | 107 | \<^descr> \<^ML>\<open>Named_Target.init\<close>~\<open>includes name thy\<close> initializes a local theory | 
| 61854 | 108 | derived from the given background theory. An empty name refers to a \<^emph>\<open>global | 
| 109 |   theory\<close> context, and a non-empty name refers to a @{command locale} or
 | |
| 110 |   @{command class} context (a fully-qualified internal name is expected here).
 | |
| 111 | This is useful for experimentation --- normally the Isar toplevel already | |
| 57181 
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
 haftmann parents: 
56420diff
changeset | 112 | takes care to initialize the local theory context. | 
| 29764 | 113 | |
| 69597 | 114 | \<^descr> \<^ML>\<open>Local_Theory.define\<close>~\<open>((b, mx), (a, rhs)) lthy\<close> defines a local | 
| 61854 | 115 | entity according to the specification that is given relatively to the | 
| 116 | current \<open>lthy\<close> context. In particular the term of the RHS may refer to | |
| 117 | earlier local entities from the auxiliary context, or hypothetical | |
| 118 | parameters from the target context. The result is the newly defined term | |
| 119 | (which is always a fixed variable with exactly the same name as specified | |
| 120 | for the LHS), together with an equational theorem that states the definition | |
| 121 | as a hypothetical fact. | |
| 29765 | 122 | |
| 61854 | 123 | Unless an explicit name binding is given for the RHS, the resulting fact | 
| 124 | will be called \<open>b_def\<close>. Any given attributes are applied to that same fact | |
| 125 | --- immediately in the auxiliary context \<^emph>\<open>and\<close> in any transformed versions | |
| 126 | stemming from target-specific policies or any later interpretations of | |
| 127 |   results from the target context (think of @{command locale} and @{command
 | |
| 128 | interpretation}, for example). This means that attributes should be usually | |
| 129 |   plain declarations such as @{attribute simp}, while non-trivial rules like
 | |
| 29765 | 130 |   @{attribute simplified} are better avoided.
 | 
| 131 | ||
| 69597 | 132 | \<^descr> \<^ML>\<open>Local_Theory.note\<close>~\<open>(a, ths) lthy\<close> is analogous to \<^ML>\<open>Local_Theory.define\<close>, but defines facts instead of terms. There is also a | 
| 133 | slightly more general variant \<^ML>\<open>Local_Theory.notes\<close> that defines several | |
| 61854 | 134 | facts (with attribute expressions) simultaneously. | 
| 29765 | 135 | |
| 61854 | 136 |   This is essentially the internal version of the @{command lemmas} command,
 | 
| 137 |   or @{command declare} if an empty name binding is given.
 | |
| 58618 | 138 | \<close> | 
| 18537 | 139 | |
| 20451 | 140 | |
| 58618 | 141 | section \<open>Morphisms and declarations \label{sec:morphisms}\<close>
 | 
| 18537 | 142 | |
| 58618 | 143 | text \<open> | 
| 52422 | 144 | %FIXME | 
| 39877 | 145 | |
| 76987 | 146 | See also \<^cite>\<open>"Chaieb-Wenzel:2007"\<close>. | 
| 58618 | 147 | \<close> | 
| 30272 | 148 | |
| 18537 | 149 | end |