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