doc-src/IsarRef/Thy/Spec.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 37096 67934c40a5f7
child 37768 e86221f3bac7
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
     1 theory Spec
     2 imports Main
     3 begin
     4 
     5 chapter {* Theory specifications *}
     6 
     7 text {*
     8   The Isabelle/Isar theory format integrates specifications and
     9   proofs, supporting interactive development with unlimited undo
    10   operation.  There is an integrated document preparation system (see
    11   \chref{ch:document-prep}), for typesetting formal developments
    12   together with informal text.  The resulting hyper-linked PDF
    13   documents can be used both for WWW presentation and printed copies.
    14 
    15   The Isar proof language (see \chref{ch:proofs}) is embedded into the
    16   theory language as a proper sub-language.  Proof mode is entered by
    17   stating some @{command theorem} or @{command lemma} at the theory
    18   level, and left again with the final conclusion (e.g.\ via @{command
    19   qed}).  Some theory specification mechanisms also require a proof,
    20   such as @{command typedef} in HOL, which demands non-emptiness of
    21   the representing sets.
    22 *}
    23 
    24 
    25 section {* Defining theories \label{sec:begin-thy} *}
    26 
    27 text {*
    28   \begin{matharray}{rcl}
    29     @{command_def "theory"} & : & @{text "toplevel \<rightarrow> theory"} \\
    30     @{command_def (global) "end"} & : & @{text "theory \<rightarrow> toplevel"} \\
    31   \end{matharray}
    32 
    33   Isabelle/Isar theories are defined via theory files, which may
    34   contain both specifications and proofs; occasionally definitional
    35   mechanisms also require some explicit proof.  The theory body may be
    36   sub-structured by means of \emph{local theory targets}, such as
    37   @{command "locale"} and @{command "class"}.
    38 
    39   The first proper command of a theory is @{command "theory"}, which
    40   indicates imports of previous theories and optional dependencies on
    41   other source files (usually in ML).  Just preceding the initial
    42   @{command "theory"} command there may be an optional @{command
    43   "header"} declaration, which is only relevant to document
    44   preparation: see also the other section markup commands in
    45   \secref{sec:markup}.
    46 
    47   A theory is concluded by a final @{command (global) "end"} command,
    48   one that does not belong to a local theory target.  No further
    49   commands may follow such a global @{command (global) "end"},
    50   although some user-interfaces might pretend that trailing input is
    51   admissible.
    52 
    53   \begin{rail}
    54     'theory' name 'imports' (name +) uses? 'begin'
    55     ;
    56 
    57     uses: 'uses' ((name | parname) +);
    58   \end{rail}
    59 
    60   \begin{description}
    61 
    62   \item @{command "theory"}~@{text "A \<IMPORTS> B\<^sub>1 \<dots> B\<^sub>n \<BEGIN>"}
    63   starts a new theory @{text A} based on the merge of existing
    64   theories @{text "B\<^sub>1 \<dots> B\<^sub>n"}.
    65   
    66   Due to the possibility to import more than one ancestor, the
    67   resulting theory structure of an Isabelle session forms a directed
    68   acyclic graph (DAG).  Isabelle's theory loader ensures that the
    69   sources contributing to the development graph are always up-to-date:
    70   changed files are automatically reloaded whenever a theory header
    71   specification is processed.
    72   
    73   The optional @{keyword_def "uses"} specification declares additional
    74   dependencies on extra files (usually ML sources).  Files will be
    75   loaded immediately (as ML), unless the name is parenthesized.  The
    76   latter case records a dependency that needs to be resolved later in
    77   the text, usually via explicit @{command_ref "use"} for ML files;
    78   other file formats require specific load commands defined by the
    79   corresponding tools or packages.
    80   
    81   \item @{command (global) "end"} concludes the current theory
    82   definition.  Note that local theory targets involve a local
    83   @{command (local) "end"}, which is clear from the nesting.
    84 
    85   \end{description}
    86 *}
    87 
    88 
    89 section {* Local theory targets \label{sec:target} *}
    90 
    91 text {*
    92   A local theory target is a context managed separately within the
    93   enclosing theory.  Contexts may introduce parameters (fixed
    94   variables) and assumptions (hypotheses).  Definitions and theorems
    95   depending on the context may be added incrementally later on.  Named
    96   contexts refer to locales (cf.\ \secref{sec:locale}) or type classes
    97   (cf.\ \secref{sec:class}); the name ``@{text "-"}'' signifies the
    98   global theory context.
    99 
   100   \begin{matharray}{rcll}
   101     @{command_def "context"} & : & @{text "theory \<rightarrow> local_theory"} \\
   102     @{command_def (local) "end"} & : & @{text "local_theory \<rightarrow> theory"} \\
   103   \end{matharray}
   104 
   105   \indexouternonterm{target}
   106   \begin{rail}
   107     'context' name 'begin'
   108     ;
   109 
   110     target: '(' 'in' name ')'
   111     ;
   112   \end{rail}
   113 
   114   \begin{description}
   115   
   116   \item @{command "context"}~@{text "c \<BEGIN>"} recommences an
   117   existing locale or class context @{text c}.  Note that locale and
   118   class definitions allow to include the @{keyword "begin"} keyword as
   119   well, in order to continue the local theory immediately after the
   120   initial specification.
   121   
   122   \item @{command (local) "end"} concludes the current local theory
   123   and continues the enclosing global theory.  Note that a global
   124   @{command (global) "end"} has a different meaning: it concludes the
   125   theory itself (\secref{sec:begin-thy}).
   126   
   127   \item @{text "("}@{keyword_def "in"}~@{text "c)"} given after any
   128   local theory command specifies an immediate target, e.g.\
   129   ``@{command "definition"}~@{text "(\<IN> c) \<dots>"}'' or ``@{command
   130   "theorem"}~@{text "(\<IN> c) \<dots>"}''.  This works both in a local or
   131   global theory context; the current target context will be suspended
   132   for this command only.  Note that ``@{text "(\<IN> -)"}'' will
   133   always produce a global result independently of the current target
   134   context.
   135 
   136   \end{description}
   137 
   138   The exact meaning of results produced within a local theory context
   139   depends on the underlying target infrastructure (locale, type class
   140   etc.).  The general idea is as follows, considering a context named
   141   @{text c} with parameter @{text x} and assumption @{text "A[x]"}.
   142   
   143   Definitions are exported by introducing a global version with
   144   additional arguments; a syntactic abbreviation links the long form
   145   with the abstract version of the target context.  For example,
   146   @{text "a \<equiv> t[x]"} becomes @{text "c.a ?x \<equiv> t[?x]"} at the theory
   147   level (for arbitrary @{text "?x"}), together with a local
   148   abbreviation @{text "c \<equiv> c.a x"} in the target context (for the
   149   fixed parameter @{text x}).
   150 
   151   Theorems are exported by discharging the assumptions and
   152   generalizing the parameters of the context.  For example, @{text "a:
   153   B[x]"} becomes @{text "c.a: A[?x] \<Longrightarrow> B[?x]"}, again for arbitrary
   154   @{text "?x"}.
   155 *}
   156 
   157 
   158 section {* Basic specification elements *}
   159 
   160 text {*
   161   \begin{matharray}{rcll}
   162     @{command_def "axiomatization"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!)\\
   163     @{command_def "definition"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   164     @{attribute_def "defn"} & : & @{text attribute} \\
   165     @{command_def "abbreviation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   166     @{command_def "print_abbrevs"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow> "} \\
   167   \end{matharray}
   168 
   169   These specification mechanisms provide a slightly more abstract view
   170   than the underlying primitives of @{command "consts"}, @{command
   171   "defs"} (see \secref{sec:consts}), and @{command "axioms"} (see
   172   \secref{sec:axms-thms}).  In particular, type-inference is commonly
   173   available, and result names need not be given.
   174 
   175   \begin{rail}
   176     'axiomatization' target? fixes? ('where' specs)?
   177     ;
   178     'definition' target? (decl 'where')? thmdecl? prop
   179     ;
   180     'abbreviation' target? mode? (decl 'where')? prop
   181     ;
   182 
   183     fixes: ((name ('::' type)? mixfix? | vars) + 'and')
   184     ;
   185     specs: (thmdecl? props + 'and')
   186     ;
   187     decl: name ('::' type)? mixfix?
   188     ;
   189   \end{rail}
   190 
   191   \begin{description}
   192   
   193   \item @{command "axiomatization"}~@{text "c\<^sub>1 \<dots> c\<^sub>m \<WHERE> \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"}
   194   introduces several constants simultaneously and states axiomatic
   195   properties for these.  The constants are marked as being specified
   196   once and for all, which prevents additional specifications being
   197   issued later on.
   198   
   199   Note that axiomatic specifications are only appropriate when
   200   declaring a new logical system; axiomatic specifications are
   201   restricted to global theory contexts.  Normal applications should
   202   only use definitional mechanisms!
   203 
   204   \item @{command "definition"}~@{text "c \<WHERE> eq"} produces an
   205   internal definition @{text "c \<equiv> t"} according to the specification
   206   given as @{text eq}, which is then turned into a proven fact.  The
   207   given proposition may deviate from internal meta-level equality
   208   according to the rewrite rules declared as @{attribute defn} by the
   209   object-logic.  This usually covers object-level equality @{text "x =
   210   y"} and equivalence @{text "A \<leftrightarrow> B"}.  End-users normally need not
   211   change the @{attribute defn} setup.
   212   
   213   Definitions may be presented with explicit arguments on the LHS, as
   214   well as additional conditions, e.g.\ @{text "f x y = t"} instead of
   215   @{text "f \<equiv> \<lambda>x y. t"} and @{text "y \<noteq> 0 \<Longrightarrow> g x y = u"} instead of an
   216   unrestricted @{text "g \<equiv> \<lambda>x y. u"}.
   217   
   218   \item @{command "abbreviation"}~@{text "c \<WHERE> eq"} introduces a
   219   syntactic constant which is associated with a certain term according
   220   to the meta-level equality @{text eq}.
   221   
   222   Abbreviations participate in the usual type-inference process, but
   223   are expanded before the logic ever sees them.  Pretty printing of
   224   terms involves higher-order rewriting with rules stemming from
   225   reverted abbreviations.  This needs some care to avoid overlapping
   226   or looping syntactic replacements!
   227   
   228   The optional @{text mode} specification restricts output to a
   229   particular print mode; using ``@{text input}'' here achieves the
   230   effect of one-way abbreviations.  The mode may also include an
   231   ``@{keyword "output"}'' qualifier that affects the concrete syntax
   232   declared for abbreviations, cf.\ @{command "syntax"} in
   233   \secref{sec:syn-trans}.
   234   
   235   \item @{command "print_abbrevs"} prints all constant abbreviations
   236   of the current context.
   237   
   238   \end{description}
   239 *}
   240 
   241 
   242 section {* Generic declarations *}
   243 
   244 text {*
   245   Arbitrary operations on the background context may be wrapped-up as
   246   generic declaration elements.  Since the underlying concept of local
   247   theories may be subject to later re-interpretation, there is an
   248   additional dependency on a morphism that tells the difference of the
   249   original declaration context wrt.\ the application context
   250   encountered later on.  A fact declaration is an important special
   251   case: it consists of a theorem which is applied to the context by
   252   means of an attribute.
   253 
   254   \begin{matharray}{rcl}
   255     @{command_def "declaration"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   256     @{command_def "declare"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   257   \end{matharray}
   258 
   259   \begin{rail}
   260     'declaration' ('(pervasive)')? target? text
   261     ;
   262     'declare' target? (thmrefs + 'and')
   263     ;
   264   \end{rail}
   265 
   266   \begin{description}
   267 
   268   \item @{command "declaration"}~@{text d} adds the declaration
   269   function @{text d} of ML type @{ML_type declaration}, to the current
   270   local theory under construction.  In later application contexts, the
   271   function is transformed according to the morphisms being involved in
   272   the interpretation hierarchy.
   273 
   274   If the @{text "(pervasive)"} option is given, the corresponding
   275   declaration is applied to all possible contexts involved, including
   276   the global background theory.
   277 
   278   \item @{command "declare"}~@{text thms} declares theorems to the
   279   current local theory context.  No theorem binding is involved here,
   280   unlike @{command "theorems"} or @{command "lemmas"} (cf.\
   281   \secref{sec:axms-thms}), so @{command "declare"} only has the effect
   282   of applying attributes as included in the theorem specification.
   283 
   284   \end{description}
   285 *}
   286 
   287 
   288 section {* Locales \label{sec:locale} *}
   289 
   290 text {*
   291   Locales are parametric named local contexts, consisting of a list of
   292   declaration elements that are modeled after the Isar proof context
   293   commands (cf.\ \secref{sec:proof-context}).
   294 *}
   295 
   296 
   297 subsection {* Locale expressions \label{sec:locale-expr} *}
   298 
   299 text {*
   300   A \emph{locale expression} denotes a structured context composed of
   301   instances of existing locales.  The context consists of a list of
   302   instances of declaration elements from the locales.  Two locale
   303   instances are equal if they are of the same locale and the
   304   parameters are instantiated with equivalent terms.  Declaration
   305   elements from equal instances are never repeated, thus avoiding
   306   duplicate declarations.
   307 
   308   \indexouternonterm{localeexpr}
   309   \begin{rail}
   310     localeexpr: (instance + '+') ('for' (fixes + 'and'))?
   311     ;
   312     instance: (qualifier ':')? nameref (posinsts | namedinsts)
   313     ;
   314     qualifier: name ('?' | '!')?
   315     ;
   316     posinsts: (term | '_')*
   317     ;
   318     namedinsts: 'where' (name '=' term + 'and')
   319     ;
   320   \end{rail}
   321 
   322   A locale instance consists of a reference to a locale and either
   323   positional or named parameter instantiations.  Identical
   324   instantiations (that is, those that instante a parameter by itself)
   325   may be omitted.  The notation `\_' enables to omit the instantiation
   326   for a parameter inside a positional instantiation.
   327 
   328   Terms in instantiations are from the context the locale expressions
   329   is declared in.  Local names may be added to this context with the
   330   optional for clause.  In addition, syntax declarations from one
   331   instance are effective when parsing subsequent instances of the same
   332   expression.
   333 
   334   Instances have an optional qualifier which applies to names in
   335   declarations.  Names include local definitions and theorem names.
   336   If present, the qualifier itself is either optional
   337   (``\texttt{?}''), which means that it may be omitted on input of the
   338   qualified name, or mandatory (``\texttt{!}'').  If neither
   339   ``\texttt{?}'' nor ``\texttt{!}'' are present, the command's default
   340   is used.  For @{command "interpretation"} and @{command "interpret"}
   341   the default is ``mandatory'', for @{command "locale"} and @{command
   342   "sublocale"} the default is ``optional''.
   343 *}
   344 
   345 
   346 subsection {* Locale declarations *}
   347 
   348 text {*
   349   \begin{matharray}{rcl}
   350     @{command_def "locale"} & : & @{text "theory \<rightarrow> local_theory"} \\
   351     @{command_def "print_locale"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   352     @{command_def "print_locales"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   353     @{method_def intro_locales} & : & @{text method} \\
   354     @{method_def unfold_locales} & : & @{text method} \\
   355   \end{matharray}
   356 
   357   \indexouternonterm{contextelem}
   358   \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes}
   359   \indexisarelem{defines}\indexisarelem{notes}
   360   \begin{rail}
   361     'locale' name ('=' locale)? 'begin'?
   362     ;
   363     'print\_locale' '!'? nameref
   364     ;
   365     locale: contextelem+ | localeexpr ('+' (contextelem+))?
   366     ;
   367     contextelem:
   368     'fixes' (fixes + 'and')
   369     | 'constrains' (name '::' type + 'and')
   370     | 'assumes' (props + 'and')
   371     | 'defines' (thmdecl? prop proppat? + 'and')
   372     | 'notes' (thmdef? thmrefs + 'and')
   373     ;
   374   \end{rail}
   375 
   376   \begin{description}
   377   
   378   \item @{command "locale"}~@{text "loc = import + body"} defines a
   379   new locale @{text loc} as a context consisting of a certain view of
   380   existing locales (@{text import}) plus some additional elements
   381   (@{text body}).  Both @{text import} and @{text body} are optional;
   382   the degenerate form @{command "locale"}~@{text loc} defines an empty
   383   locale, which may still be useful to collect declarations of facts
   384   later on.  Type-inference on locale expressions automatically takes
   385   care of the most general typing that the combined context elements
   386   may acquire.
   387 
   388   The @{text import} consists of a structured locale expression; see
   389   \secref{sec:proof-context} above.  Its for clause defines the local
   390   parameters of the @{text import}.  In addition, locale parameters
   391   whose instantance is omitted automatically extend the (possibly
   392   empty) for clause: they are inserted at its beginning.  This means
   393   that these parameters may be referred to from within the expression
   394   and also in the subsequent context elements and provides a
   395   notational convenience for the inheritance of parameters in locale
   396   declarations.
   397 
   398   The @{text body} consists of context elements.
   399 
   400   \begin{description}
   401 
   402   \item @{element "fixes"}~@{text "x :: \<tau> (mx)"} declares a local
   403   parameter of type @{text \<tau>} and mixfix annotation @{text mx} (both
   404   are optional).  The special syntax declaration ``@{text
   405   "(\<STRUCTURE>)"}'' means that @{text x} may be referenced
   406   implicitly in this context.
   407 
   408   \item @{element "constrains"}~@{text "x :: \<tau>"} introduces a type
   409   constraint @{text \<tau>} on the local parameter @{text x}.  This
   410   element is deprecated.  The type constaint should be introduced in
   411   the for clause or the relevant @{element "fixes"} element.
   412 
   413   \item @{element "assumes"}~@{text "a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"}
   414   introduces local premises, similar to @{command "assume"} within a
   415   proof (cf.\ \secref{sec:proof-context}).
   416 
   417   \item @{element "defines"}~@{text "a: x \<equiv> t"} defines a previously
   418   declared parameter.  This is similar to @{command "def"} within a
   419   proof (cf.\ \secref{sec:proof-context}), but @{element "defines"}
   420   takes an equational proposition instead of variable-term pair.  The
   421   left-hand side of the equation may have additional arguments, e.g.\
   422   ``@{element "defines"}~@{text "f x\<^sub>1 \<dots> x\<^sub>n \<equiv> t"}''.
   423 
   424   \item @{element "notes"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"}
   425   reconsiders facts within a local context.  Most notably, this may
   426   include arbitrary declarations in any attribute specifications
   427   included here, e.g.\ a local @{attribute simp} rule.
   428 
   429   The initial @{text import} specification of a locale expression
   430   maintains a dynamic relation to the locales being referenced
   431   (benefiting from any later fact declarations in the obvious manner).
   432 
   433   \end{description}
   434   
   435   Note that ``@{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"}'' patterns given
   436   in the syntax of @{element "assumes"} and @{element "defines"} above
   437   are illegal in locale definitions.  In the long goal format of
   438   \secref{sec:goals}, term bindings may be included as expected,
   439   though.
   440   
   441   \medskip Locale specifications are ``closed up'' by
   442   turning the given text into a predicate definition @{text
   443   loc_axioms} and deriving the original assumptions as local lemmas
   444   (modulo local definitions).  The predicate statement covers only the
   445   newly specified assumptions, omitting the content of included locale
   446   expressions.  The full cumulative view is only provided on export,
   447   involving another predicate @{text loc} that refers to the complete
   448   specification text.
   449   
   450   In any case, the predicate arguments are those locale parameters
   451   that actually occur in the respective piece of text.  Also note that
   452   these predicates operate at the meta-level in theory, but the locale
   453   packages attempts to internalize statements according to the
   454   object-logic setup (e.g.\ replacing @{text \<And>} by @{text \<forall>}, and
   455   @{text "\<Longrightarrow>"} by @{text "\<longrightarrow>"} in HOL; see also
   456   \secref{sec:object-logic}).  Separate introduction rules @{text
   457   loc_axioms.intro} and @{text loc.intro} are provided as well.
   458   
   459   \item @{command "print_locale"}~@{text "locale"} prints the
   460   contents of the named locale.  The command omits @{element "notes"}
   461   elements by default.  Use @{command "print_locale"}@{text "!"} to
   462   have them included.
   463 
   464   \item @{command "print_locales"} prints the names of all locales
   465   of the current theory.
   466 
   467   \item @{method intro_locales} and @{method unfold_locales}
   468   repeatedly expand all introduction rules of locale predicates of the
   469   theory.  While @{method intro_locales} only applies the @{text
   470   loc.intro} introduction rules and therefore does not decend to
   471   assumptions, @{method unfold_locales} is more aggressive and applies
   472   @{text loc_axioms.intro} as well.  Both methods are aware of locale
   473   specifications entailed by the context, both from target statements,
   474   and from interpretations (see below).  New goals that are entailed
   475   by the current context are discharged automatically.
   476 
   477   \end{description}
   478 *}
   479 
   480 
   481 subsection {* Locale interpretations *}
   482 
   483 text {*
   484   Locale expressions may be instantiated, and the instantiated facts
   485   added to the current context.  This requires a proof of the
   486   instantiated specification and is called \emph{locale
   487   interpretation}.  Interpretation is possible in locales @{command
   488   "sublocale"}, theories (command @{command "interpretation"}) and
   489   also within a proof body (command @{command "interpret"}).
   490 
   491   \begin{matharray}{rcl}
   492     @{command_def "sublocale"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   493     @{command_def "interpretation"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   494     @{command_def "interpret"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
   495     @{command_def "print_interps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   496   \end{matharray}
   497 
   498   \indexouternonterm{interp}
   499   \begin{rail}
   500     'sublocale' nameref ('<' | subseteq) localeexpr
   501     ;
   502     'interpretation' localeepxr equations?
   503     ;
   504     'interpret' localeexpr
   505     ;
   506     'print\_interps' nameref
   507     ;
   508     equations: 'where' (thmdecl? prop + 'and')
   509     ;
   510   \end{rail}
   511 
   512   \begin{description}
   513 
   514   \item @{command "sublocale"}~@{text "name \<subseteq> expr"}
   515   interprets @{text expr} in the locale @{text name}.  A proof that
   516   the specification of @{text name} implies the specification of
   517   @{text expr} is required.  As in the localized version of the
   518   theorem command, the proof is in the context of @{text name}.  After
   519   the proof obligation has been dischared, the facts of @{text expr}
   520   become part of locale @{text name} as \emph{derived} context
   521   elements and are available when the context @{text name} is
   522   subsequently entered.  Note that, like import, this is dynamic:
   523   facts added to a locale part of @{text expr} after interpretation
   524   become also available in @{text name}.
   525 
   526   Only specification fragments of @{text expr} that are not already
   527   part of @{text name} (be it imported, derived or a derived fragment
   528   of the import) are considered in this process.  This enables
   529   circular interpretations to the extent that no infinite chains are
   530   generated in the locale hierarchy.
   531 
   532   If interpretations of @{text name} exist in the current theory, the
   533   command adds interpretations for @{text expr} as well, with the same
   534   qualifier, although only for fragments of @{text expr} that are not
   535   interpreted in the theory already.
   536 
   537   \item @{command "interpretation"}~@{text "expr \<WHERE> eqns"}
   538   interprets @{text expr} in the theory.  The command generates proof
   539   obligations for the instantiated specifications (assumes and defines
   540   elements).  Once these are discharged by the user, instantiated
   541   facts are added to the theory in a post-processing phase.
   542 
   543   Additional equations, which are unfolded during
   544   post-processing, may be given after the keyword @{keyword "where"}.
   545   This is useful for interpreting concepts introduced through
   546   definition specification elements.  The equations must be proved.
   547 
   548   The command is aware of interpretations already active in the
   549   theory, but does not simplify the goal automatically.  In order to
   550   simplify the proof obligations use methods @{method intro_locales}
   551   or @{method unfold_locales}.  Post-processing is not applied to
   552   facts of interpretations that are already active.  This avoids
   553   duplication of interpreted facts, in particular.  Note that, in the
   554   case of a locale with import, parts of the interpretation may
   555   already be active.  The command will only process facts for new
   556   parts.
   557 
   558   Adding facts to locales has the effect of adding interpreted facts
   559   to the theory for all active interpretations also.  That is,
   560   interpretations dynamically participate in any facts added to
   561   locales.
   562 
   563   \item @{command "interpret"}~@{text "expr"}
   564   interprets @{text expr} in the proof context and is otherwise
   565   similar to interpretation in theories.
   566 
   567   \item @{command "print_interps"}~@{text "locale"} lists all
   568   interpretations of @{text "locale"} in the current theory, including
   569   those due to a combination of an @{command "interpretation"} and
   570   one or several @{command "sublocale"} declarations.
   571 
   572   \end{description}
   573 
   574   \begin{warn}
   575     Since attributes are applied to interpreted theorems,
   576     interpretation may modify the context of common proof tools, e.g.\
   577     the Simplifier or Classical Reasoner.  As the behavior of such
   578     tools is \emph{not} stable under interpretation morphisms, manual
   579     declarations might have to be added to the target context of the
   580     interpretation to revert such declarations.
   581   \end{warn}
   582 
   583   \begin{warn}
   584     An interpretation in a theory may subsume previous
   585     interpretations.  This happens if the same specification fragment
   586     is interpreted twice and the instantiation of the second
   587     interpretation is more general than the interpretation of the
   588     first.  The locale package does not attempt to remove subsumed
   589     interpretations.
   590   \end{warn}
   591 *}
   592 
   593 
   594 section {* Classes \label{sec:class} *}
   595 
   596 text {*
   597   A class is a particular locale with \emph{exactly one} type variable
   598   @{text \<alpha>}.  Beyond the underlying locale, a corresponding type class
   599   is established which is interpreted logically as axiomatic type
   600   class \cite{Wenzel:1997:TPHOL} whose logical content are the
   601   assumptions of the locale.  Thus, classes provide the full
   602   generality of locales combined with the commodity of type classes
   603   (notably type-inference).  See \cite{isabelle-classes} for a short
   604   tutorial.
   605 
   606   \begin{matharray}{rcl}
   607     @{command_def "class"} & : & @{text "theory \<rightarrow> local_theory"} \\
   608     @{command_def "instantiation"} & : & @{text "theory \<rightarrow> local_theory"} \\
   609     @{command_def "instance"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   610     @{command_def "instance"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   611     @{command_def "subclass"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   612     @{command_def "print_classes"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   613     @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   614     @{method_def intro_classes} & : & @{text method} \\
   615   \end{matharray}
   616 
   617   \begin{rail}
   618     'class' name '=' ((superclassexpr '+' (contextelem+)) | superclassexpr | (contextelem+)) \\
   619       'begin'?
   620     ;
   621     'instantiation' (nameref + 'and') '::' arity 'begin'
   622     ;
   623     'instance'
   624     ;
   625     'instance' (nameref + 'and') '::' arity
   626     ;
   627     'subclass' target? nameref
   628     ;
   629     'instance' nameref ('<' | subseteq) nameref
   630     ;
   631     'print\_classes'
   632     ;
   633     'class\_deps'
   634     ;
   635 
   636     superclassexpr: nameref | (nameref '+' superclassexpr)
   637     ;
   638   \end{rail}
   639 
   640   \begin{description}
   641 
   642   \item @{command "class"}~@{text "c = superclasses + body"} defines
   643   a new class @{text c}, inheriting from @{text superclasses}.  This
   644   introduces a locale @{text c} with import of all locales @{text
   645   superclasses}.
   646 
   647   Any @{element "fixes"} in @{text body} are lifted to the global
   648   theory level (\emph{class operations} @{text "f\<^sub>1, \<dots>,
   649   f\<^sub>n"} of class @{text c}), mapping the local type parameter
   650   @{text \<alpha>} to a schematic type variable @{text "?\<alpha> :: c"}.
   651 
   652   Likewise, @{element "assumes"} in @{text body} are also lifted,
   653   mapping each local parameter @{text "f :: \<tau>[\<alpha>]"} to its
   654   corresponding global constant @{text "f :: \<tau>[?\<alpha> :: c]"}.  The
   655   corresponding introduction rule is provided as @{text
   656   c_class_axioms.intro}.  This rule should be rarely needed directly
   657   --- the @{method intro_classes} method takes care of the details of
   658   class membership proofs.
   659 
   660   \item @{command "instantiation"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s
   661   \<BEGIN>"} opens a theory target (cf.\ \secref{sec:target}) which
   662   allows to specify class operations @{text "f\<^sub>1, \<dots>, f\<^sub>n"} corresponding
   663   to sort @{text s} at the particular type instance @{text "(\<alpha>\<^sub>1 :: s\<^sub>1,
   664   \<dots>, \<alpha>\<^sub>n :: s\<^sub>n) t"}.  A plain @{command "instance"} command in the
   665   target body poses a goal stating these type arities.  The target is
   666   concluded by an @{command_ref (local) "end"} command.
   667 
   668   Note that a list of simultaneous type constructors may be given;
   669   this corresponds nicely to mutually recursive type definitions, e.g.\
   670   in Isabelle/HOL.
   671 
   672   \item @{command "instance"} in an instantiation target body sets
   673   up a goal stating the type arities claimed at the opening @{command
   674   "instantiation"}.  The proof would usually proceed by @{method
   675   intro_classes}, and then establish the characteristic theorems of
   676   the type classes involved.  After finishing the proof, the
   677   background theory will be augmented by the proven type arities.
   678 
   679   On the theory level, @{command "instance"}~@{text "t :: (s\<^sub>1, \<dots>,
   680   s\<^sub>n)s"} provides a convenient way to instantiate a type class with no
   681   need to specify operations: one can continue with the
   682   instantiation proof immediately.
   683 
   684   \item @{command "subclass"}~@{text c} in a class context for class
   685   @{text d} sets up a goal stating that class @{text c} is logically
   686   contained in class @{text d}.  After finishing the proof, class
   687   @{text d} is proven to be subclass @{text c} and the locale @{text
   688   c} is interpreted into @{text d} simultaneously.
   689 
   690   A weakend form of this is available through a further variant of
   691   @{command instance}:  @{command instance}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} opens
   692   a proof that class @{text "c\<^isub>2"} implies @{text "c\<^isub>1"} without reference
   693   to the underlying locales;  this is useful if the properties to prove
   694   the logical connection are not sufficent on the locale level but on
   695   the theory level.
   696 
   697   \item @{command "print_classes"} prints all classes in the current
   698   theory.
   699 
   700   \item @{command "class_deps"} visualizes all classes and their
   701   subclass relations as a Hasse diagram.
   702 
   703   \item @{method intro_classes} repeatedly expands all class
   704   introduction rules of this theory.  Note that this method usually
   705   needs not be named explicitly, as it is already included in the
   706   default proof step (e.g.\ of @{command "proof"}).  In particular,
   707   instantiation of trivial (syntactic) classes may be performed by a
   708   single ``@{command ".."}'' proof step.
   709 
   710   \end{description}
   711 *}
   712 
   713 
   714 subsection {* The class target *}
   715 
   716 text {*
   717   %FIXME check
   718 
   719   A named context may refer to a locale (cf.\ \secref{sec:target}).
   720   If this locale is also a class @{text c}, apart from the common
   721   locale target behaviour the following happens.
   722 
   723   \begin{itemize}
   724 
   725   \item Local constant declarations @{text "g[\<alpha>]"} referring to the
   726   local type parameter @{text \<alpha>} and local parameters @{text "f[\<alpha>]"}
   727   are accompanied by theory-level constants @{text "g[?\<alpha> :: c]"}
   728   referring to theory-level class operations @{text "f[?\<alpha> :: c]"}.
   729 
   730   \item Local theorem bindings are lifted as are assumptions.
   731 
   732   \item Local syntax refers to local operations @{text "g[\<alpha>]"} and
   733   global operations @{text "g[?\<alpha> :: c]"} uniformly.  Type inference
   734   resolves ambiguities.  In rare cases, manual type annotations are
   735   needed.
   736   
   737   \end{itemize}
   738 *}
   739 
   740 
   741 section {* Unrestricted overloading *}
   742 
   743 text {*
   744   Isabelle/Pure's definitional schemes support certain forms of
   745   overloading (see \secref{sec:consts}).  Overloading means that a
   746   constant being declared as @{text "c :: \<alpha> decl"} may be
   747   defined separately on type instances
   748   @{text "c :: (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n) t decl"}
   749   for each type constructor @{text t}.  At most occassions
   750   overloading will be used in a Haskell-like fashion together with
   751   type classes by means of @{command "instantiation"} (see
   752   \secref{sec:class}).  Sometimes low-level overloading is desirable.
   753   The @{command "overloading"} target provides a convenient view for
   754   end-users.
   755 
   756   \begin{matharray}{rcl}
   757     @{command_def "overloading"} & : & @{text "theory \<rightarrow> local_theory"} \\
   758   \end{matharray}
   759 
   760   \begin{rail}
   761     'overloading' \\
   762     ( string ( '==' | equiv ) term ( '(' 'unchecked' ')' )? + ) 'begin'
   763   \end{rail}
   764 
   765   \begin{description}
   766 
   767   \item @{command "overloading"}~@{text "x\<^sub>1 \<equiv> c\<^sub>1 :: \<tau>\<^sub>1 \<AND> \<dots> x\<^sub>n \<equiv> c\<^sub>n :: \<tau>\<^sub>n \<BEGIN>"}
   768   opens a theory target (cf.\ \secref{sec:target}) which allows to
   769   specify constants with overloaded definitions.  These are identified
   770   by an explicitly given mapping from variable names @{text "x\<^sub>i"} to
   771   constants @{text "c\<^sub>i"} at particular type instances.  The
   772   definitions themselves are established using common specification
   773   tools, using the names @{text "x\<^sub>i"} as reference to the
   774   corresponding constants.  The target is concluded by @{command
   775   (local) "end"}.
   776 
   777   A @{text "(unchecked)"} option disables global dependency checks for
   778   the corresponding definition, which is occasionally useful for
   779   exotic overloading (see \secref{sec:consts} for a precise description).
   780   It is at the discretion of the user to avoid
   781   malformed theory specifications!
   782 
   783   \end{description}
   784 *}
   785 
   786 
   787 section {* Incorporating ML code \label{sec:ML} *}
   788 
   789 text {*
   790   \begin{matharray}{rcl}
   791     @{command_def "use"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   792     @{command_def "ML"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   793     @{command_def "ML_prf"} & : & @{text "proof \<rightarrow> proof"} \\
   794     @{command_def "ML_val"} & : & @{text "any \<rightarrow>"} \\
   795     @{command_def "ML_command"} & : & @{text "any \<rightarrow>"} \\
   796     @{command_def "setup"} & : & @{text "theory \<rightarrow> theory"} \\
   797     @{command_def "local_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   798     @{command_def "attribute_setup"} & : & @{text "theory \<rightarrow> theory"} \\
   799   \end{matharray}
   800 
   801   \begin{mldecls}
   802     @{index_ML bind_thms: "string * thm list -> unit"} \\
   803     @{index_ML bind_thm: "string * thm -> unit"} \\
   804   \end{mldecls}
   805 
   806   \begin{rail}
   807     'use' name
   808     ;
   809     ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup' | 'local\_setup') text
   810     ;
   811     'attribute\_setup' name '=' text text
   812     ;
   813   \end{rail}
   814 
   815   \begin{description}
   816 
   817   \item @{command "use"}~@{text "file"} reads and executes ML
   818   commands from @{text "file"}.  The current theory context is passed
   819   down to the ML toplevel and may be modified, using @{ML
   820   "Context.>>"} or derived ML commands.  The file name is checked with
   821   the @{keyword_ref "uses"} dependency declaration given in the theory
   822   header (see also \secref{sec:begin-thy}).
   823 
   824   Top-level ML bindings are stored within the (global or local) theory
   825   context.
   826   
   827   \item @{command "ML"}~@{text "text"} is similar to @{command "use"},
   828   but executes ML commands directly from the given @{text "text"}.
   829   Top-level ML bindings are stored within the (global or local) theory
   830   context.
   831 
   832   \item @{command "ML_prf"} is analogous to @{command "ML"} but works
   833   within a proof context.
   834 
   835   Top-level ML bindings are stored within the proof context in a
   836   purely sequential fashion, disregarding the nested proof structure.
   837   ML bindings introduced by @{command "ML_prf"} are discarded at the
   838   end of the proof.
   839 
   840   \item @{command "ML_val"} and @{command "ML_command"} are diagnostic
   841   versions of @{command "ML"}, which means that the context may not be
   842   updated.  @{command "ML_val"} echos the bindings produced at the ML
   843   toplevel, but @{command "ML_command"} is silent.
   844   
   845   \item @{command "setup"}~@{text "text"} changes the current theory
   846   context by applying @{text "text"}, which refers to an ML expression
   847   of type @{ML_type "theory -> theory"}.  This enables to initialize
   848   any object-logic specific tools and packages written in ML, for
   849   example.
   850 
   851   \item @{command "local_setup"} is similar to @{command "setup"} for
   852   a local theory context, and an ML expression of type @{ML_type
   853   "local_theory -> local_theory"}.  This allows to
   854   invoke local theory specification packages without going through
   855   concrete outer syntax, for example.
   856 
   857   \item @{command "attribute_setup"}~@{text "name = text description"}
   858   defines an attribute in the current theory.  The given @{text
   859   "text"} has to be an ML expression of type
   860   @{ML_type "attribute context_parser"}, cf.\ basic parsers defined in
   861   structure @{ML_struct Args} and @{ML_struct Attrib}.
   862 
   863   In principle, attributes can operate both on a given theorem and the
   864   implicit context, although in practice only one is modified and the
   865   other serves as parameter.  Here are examples for these two cases:
   866 
   867   \end{description}
   868 *}
   869 
   870     attribute_setup my_rule = {*
   871       Attrib.thms >> (fn ths =>
   872         Thm.rule_attribute (fn context: Context.generic => fn th: thm =>
   873           let val th' = th OF ths
   874           in th' end)) *}  "my rule"
   875 
   876     attribute_setup my_declaration = {*
   877       Attrib.thms >> (fn ths =>
   878         Thm.declaration_attribute (fn th: thm => fn context: Context.generic =>
   879           let val context' = context
   880           in context' end)) *}  "my declaration"
   881 
   882 text {*
   883   \begin{description}
   884 
   885   \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
   886   theorems produced in ML both in the theory context and the ML
   887   toplevel, associating it with the provided name.  Theorems are put
   888   into a global ``standard'' format before being stored.
   889 
   890   \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a
   891   singleton theorem.
   892   
   893   \end{description}
   894 *}
   895 
   896 
   897 section {* Primitive specification elements *}
   898 
   899 subsection {* Type classes and sorts \label{sec:classes} *}
   900 
   901 text {*
   902   \begin{matharray}{rcll}
   903     @{command_def "classes"} & : & @{text "theory \<rightarrow> theory"} \\
   904     @{command_def "classrel"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   905     @{command_def "default_sort"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   906     @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   907   \end{matharray}
   908 
   909   \begin{rail}
   910     'classes' (classdecl +)
   911     ;
   912     'classrel' (nameref ('<' | subseteq) nameref + 'and')
   913     ;
   914     'default\_sort' sort
   915     ;
   916   \end{rail}
   917 
   918   \begin{description}
   919 
   920   \item @{command "classes"}~@{text "c \<subseteq> c\<^sub>1, \<dots>, c\<^sub>n"} declares class
   921   @{text c} to be a subclass of existing classes @{text "c\<^sub>1, \<dots>, c\<^sub>n"}.
   922   Isabelle implicitly maintains the transitive closure of the class
   923   hierarchy.  Cyclic class structures are not permitted.
   924 
   925   \item @{command "classrel"}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} states subclass
   926   relations between existing classes @{text "c\<^sub>1"} and @{text "c\<^sub>2"}.
   927   This is done axiomatically!  The @{command_ref "instance"}
   928   and @{command_ref "subclass"} command
   929   (see \secref{sec:class}) provide a way to introduce proven class
   930   relations.
   931 
   932   \item @{command "default_sort"}~@{text s} makes sort @{text s} the
   933   new default sort for any type variable that is given explicitly in
   934   the text, but lacks a sort constraint (wrt.\ the current context).
   935   Type variables generated by type inference are not affected.
   936 
   937   Usually the default sort is only changed when defining a new
   938   object-logic.  For example, the default sort in Isabelle/HOL is
   939   @{text type}, the class of all HOL types.  %FIXME sort antiq?
   940 
   941   When merging theories, the default sorts of the parents are
   942   logically intersected, i.e.\ the representations as lists of classes
   943   are joined.
   944 
   945   \item @{command "class_deps"} visualizes the subclass relation,
   946   using Isabelle's graph browser tool (see also \cite{isabelle-sys}).
   947 
   948   \end{description}
   949 *}
   950 
   951 
   952 subsection {* Types and type abbreviations \label{sec:types-pure} *}
   953 
   954 text {*
   955   \begin{matharray}{rcll}
   956     @{command_def "types"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   957     @{command_def "typedecl"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   958     @{command_def "arities"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   959   \end{matharray}
   960 
   961   \begin{rail}
   962     'types' (typespec '=' type mixfix? +)
   963     ;
   964     'typedecl' typespec mixfix?
   965     ;
   966     'arities' (nameref '::' arity +)
   967     ;
   968   \end{rail}
   969 
   970   \begin{description}
   971 
   972   \item @{command "types"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = \<tau>"} introduces a
   973   \emph{type synonym} @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} for the existing type
   974   @{text "\<tau>"}.  Unlike actual type definitions, as are available in
   975   Isabelle/HOL for example, type synonyms are merely syntactic
   976   abbreviations without any logical significance.  Internally, type
   977   synonyms are fully expanded.
   978   
   979   \item @{command "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} declares a new
   980   type constructor @{text t}.  If the object-logic defines a base sort
   981   @{text s}, then the constructor is declared to operate on that, via
   982   the axiomatic specification @{command arities}~@{text "t :: (s, \<dots>,
   983   s)s"}.
   984 
   985   \item @{command "arities"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s"} augments
   986   Isabelle's order-sorted signature of types by new type constructor
   987   arities.  This is done axiomatically!  The @{command_ref "instantiation"}
   988   target (see \secref{sec:class}) provides a way to introduce
   989   proven type arities.
   990 
   991   \end{description}
   992 *}
   993 
   994 
   995 subsection {* Co-regularity of type classes and arities *}
   996 
   997 text {* The class relation together with the collection of
   998   type-constructor arities must obey the principle of
   999   \emph{co-regularity} as defined below.
  1000 
  1001   \medskip For the subsequent formulation of co-regularity we assume
  1002   that the class relation is closed by transitivity and reflexivity.
  1003   Moreover the collection of arities @{text "t :: (\<^vec>s)c"} is
  1004   completed such that @{text "t :: (\<^vec>s)c"} and @{text "c \<subseteq> c'"}
  1005   implies @{text "t :: (\<^vec>s)c'"} for all such declarations.
  1006 
  1007   Treating sorts as finite sets of classes (meaning the intersection),
  1008   the class relation @{text "c\<^sub>1 \<subseteq> c\<^sub>2"} is extended to sorts as
  1009   follows:
  1010   \[
  1011     @{text "s\<^sub>1 \<subseteq> s\<^sub>2 \<equiv> \<forall>c\<^sub>2 \<in> s\<^sub>2. \<exists>c\<^sub>1 \<in> s\<^sub>1. c\<^sub>1 \<subseteq> c\<^sub>2"}
  1012   \]
  1013 
  1014   This relation on sorts is further extended to tuples of sorts (of
  1015   the same length) in the component-wise way.
  1016 
  1017   \smallskip Co-regularity of the class relation together with the
  1018   arities relation means:
  1019   \[
  1020     @{text "t :: (\<^vec>s\<^sub>1)c\<^sub>1 \<Longrightarrow> t :: (\<^vec>s\<^sub>2)c\<^sub>2 \<Longrightarrow> c\<^sub>1 \<subseteq> c\<^sub>2 \<Longrightarrow> \<^vec>s\<^sub>1 \<subseteq> \<^vec>s\<^sub>2"}
  1021   \]
  1022   \noindent for all such arities.  In other words, whenever the result
  1023   classes of some type-constructor arities are related, then the
  1024   argument sorts need to be related in the same way.
  1025 
  1026   \medskip Co-regularity is a very fundamental property of the
  1027   order-sorted algebra of types.  For example, it entails principle
  1028   types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}.
  1029 *}
  1030 
  1031 
  1032 subsection {* Constants and definitions \label{sec:consts} *}
  1033 
  1034 text {*
  1035   Definitions essentially express abbreviations within the logic.  The
  1036   simplest form of a definition is @{text "c :: \<sigma> \<equiv> t"}, where @{text
  1037   c} is a newly declared constant.  Isabelle also allows derived forms
  1038   where the arguments of @{text c} appear on the left, abbreviating a
  1039   prefix of @{text \<lambda>}-abstractions, e.g.\ @{text "c \<equiv> \<lambda>x y. t"} may be
  1040   written more conveniently as @{text "c x y \<equiv> t"}.  Moreover,
  1041   definitions may be weakened by adding arbitrary pre-conditions:
  1042   @{text "A \<Longrightarrow> c x y \<equiv> t"}.
  1043 
  1044   \medskip The built-in well-formedness conditions for definitional
  1045   specifications are:
  1046 
  1047   \begin{itemize}
  1048 
  1049   \item Arguments (on the left-hand side) must be distinct variables.
  1050 
  1051   \item All variables on the right-hand side must also appear on the
  1052   left-hand side.
  1053 
  1054   \item All type variables on the right-hand side must also appear on
  1055   the left-hand side; this prohibits @{text "0 :: nat \<equiv> length ([] ::
  1056   \<alpha> list)"} for example.
  1057 
  1058   \item The definition must not be recursive.  Most object-logics
  1059   provide definitional principles that can be used to express
  1060   recursion safely.
  1061 
  1062   \end{itemize}
  1063 
  1064   The right-hand side of overloaded definitions may mention overloaded constants
  1065   recursively at type instances corresponding to the immediate
  1066   argument types @{text "\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n"}.  Incomplete
  1067   specification patterns impose global constraints on all occurrences,
  1068   e.g.\ @{text "d :: \<alpha> \<times> \<alpha>"} on the left-hand side means that all
  1069   corresponding occurrences on some right-hand side need to be an
  1070   instance of this, general @{text "d :: \<alpha> \<times> \<beta>"} will be disallowed.
  1071 
  1072   \begin{matharray}{rcl}
  1073     @{command_def "consts"} & : & @{text "theory \<rightarrow> theory"} \\
  1074     @{command_def "defs"} & : & @{text "theory \<rightarrow> theory"} \\
  1075     @{command_def "constdefs"} & : & @{text "theory \<rightarrow> theory"} \\
  1076   \end{matharray}
  1077 
  1078   \begin{rail}
  1079     'consts' ((name '::' type mixfix?) +)
  1080     ;
  1081     'defs' ('(' 'unchecked'? 'overloaded'? ')')? \\ (axmdecl prop +)
  1082     ;
  1083   \end{rail}
  1084 
  1085   \begin{rail}
  1086     'constdefs' structs? (constdecl? constdef +)
  1087     ;
  1088 
  1089     structs: '(' 'structure' (vars + 'and') ')'
  1090     ;
  1091     constdecl:  ((name '::' type mixfix | name '::' type | name mixfix) 'where'?) | name 'where'
  1092     ;
  1093     constdef: thmdecl? prop
  1094     ;
  1095   \end{rail}
  1096 
  1097   \begin{description}
  1098 
  1099   \item @{command "consts"}~@{text "c :: \<sigma>"} declares constant @{text
  1100   c} to have any instance of type scheme @{text \<sigma>}.  The optional
  1101   mixfix annotations may attach concrete syntax to the constants
  1102   declared.
  1103   
  1104   \item @{command "defs"}~@{text "name: eqn"} introduces @{text eqn}
  1105   as a definitional axiom for some existing constant.
  1106   
  1107   The @{text "(unchecked)"} option disables global dependency checks
  1108   for this definition, which is occasionally useful for exotic
  1109   overloading.  It is at the discretion of the user to avoid malformed
  1110   theory specifications!
  1111   
  1112   The @{text "(overloaded)"} option declares definitions to be
  1113   potentially overloaded.  Unless this option is given, a warning
  1114   message would be issued for any definitional equation with a more
  1115   special type than that of the corresponding constant declaration.
  1116   
  1117   \item @{command "constdefs"} combines constant declarations and
  1118   definitions, with type-inference taking care of the most general
  1119   typing of the given specification (the optional type constraint may
  1120   refer to type-inference dummies ``@{text _}'' as usual).  The
  1121   resulting type declaration needs to agree with that of the
  1122   specification; overloading is \emph{not} supported here!
  1123   
  1124   The constant name may be omitted altogether, if neither type nor
  1125   syntax declarations are given.  The canonical name of the
  1126   definitional axiom for constant @{text c} will be @{text c_def},
  1127   unless specified otherwise.  Also note that the given list of
  1128   specifications is processed in a strictly sequential manner, with
  1129   type-checking being performed independently.
  1130   
  1131   An optional initial context of @{text "(structure)"} declarations
  1132   admits use of indexed syntax, using the special symbol @{verbatim
  1133   "\<index>"} (printed as ``@{text "\<index>"}'').  The latter concept is
  1134   particularly useful with locales (see also \secref{sec:locale}).
  1135 
  1136   \end{description}
  1137 *}
  1138 
  1139 
  1140 section {* Axioms and theorems \label{sec:axms-thms} *}
  1141 
  1142 text {*
  1143   \begin{matharray}{rcll}
  1144     @{command_def "axioms"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
  1145     @{command_def "lemmas"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1146     @{command_def "theorems"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1147   \end{matharray}
  1148 
  1149   \begin{rail}
  1150     'axioms' (axmdecl prop +)
  1151     ;
  1152     ('lemmas' | 'theorems') target? (thmdef? thmrefs + 'and')
  1153     ;
  1154   \end{rail}
  1155 
  1156   \begin{description}
  1157   
  1158   \item @{command "axioms"}~@{text "a: \<phi>"} introduces arbitrary
  1159   statements as axioms of the meta-logic.  In fact, axioms are
  1160   ``axiomatic theorems'', and may be referred later just as any other
  1161   theorem.
  1162   
  1163   Axioms are usually only introduced when declaring new logical
  1164   systems.  Everyday work is typically done the hard way, with proper
  1165   definitions and proven theorems.
  1166   
  1167   \item @{command "lemmas"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"} retrieves and stores
  1168   existing facts in the theory context, or the specified target
  1169   context (see also \secref{sec:target}).  Typical applications would
  1170   also involve attributes, to declare Simplifier rules, for example.
  1171   
  1172   \item @{command "theorems"} is essentially the same as @{command
  1173   "lemmas"}, but marks the result as a different kind of facts.
  1174 
  1175   \end{description}
  1176 *}
  1177 
  1178 
  1179 section {* Oracles *}
  1180 
  1181 text {* Oracles allow Isabelle to take advantage of external reasoners
  1182   such as arithmetic decision procedures, model checkers, fast
  1183   tautology checkers or computer algebra systems.  Invoked as an
  1184   oracle, an external reasoner can create arbitrary Isabelle theorems.
  1185 
  1186   It is the responsibility of the user to ensure that the external
  1187   reasoner is as trustworthy as the application requires.  Another
  1188   typical source of errors is the linkup between Isabelle and the
  1189   external tool, not just its concrete implementation, but also the
  1190   required translation between two different logical environments.
  1191 
  1192   Isabelle merely guarantees well-formedness of the propositions being
  1193   asserted, and records within the internal derivation object how
  1194   presumed theorems depend on unproven suppositions.
  1195 
  1196   \begin{matharray}{rcl}
  1197     @{command_def "oracle"} & : & @{text "theory \<rightarrow> theory"} \\
  1198   \end{matharray}
  1199 
  1200   \begin{rail}
  1201     'oracle' name '=' text
  1202     ;
  1203   \end{rail}
  1204 
  1205   \begin{description}
  1206 
  1207   \item @{command "oracle"}~@{text "name = text"} turns the given ML
  1208   expression @{text "text"} of type @{ML_text "'a -> cterm"} into an
  1209   ML function of type @{ML_text "'a -> thm"}, which is bound to the
  1210   global identifier @{ML_text name}.  This acts like an infinitary
  1211   specification of axioms!  Invoking the oracle only works within the
  1212   scope of the resulting theory.
  1213 
  1214   \end{description}
  1215 
  1216   See @{"file" "~~/src/FOL/ex/Iff_Oracle.thy"} for a worked example of
  1217   defining a new primitive rule as oracle, and turning it into a proof
  1218   method.
  1219 *}
  1220 
  1221 
  1222 section {* Name spaces *}
  1223 
  1224 text {*
  1225   \begin{matharray}{rcl}
  1226     @{command_def "global"} & : & @{text "theory \<rightarrow> theory"} \\
  1227     @{command_def "local"} & : & @{text "theory \<rightarrow> theory"} \\
  1228     @{command_def "hide_class"} & : & @{text "theory \<rightarrow> theory"} \\
  1229     @{command_def "hide_type"} & : & @{text "theory \<rightarrow> theory"} \\
  1230     @{command_def "hide_const"} & : & @{text "theory \<rightarrow> theory"} \\
  1231     @{command_def "hide_fact"} & : & @{text "theory \<rightarrow> theory"} \\
  1232   \end{matharray}
  1233 
  1234   \begin{rail}
  1235     ( 'hide\_class' | 'hide\_type' | 'hide\_const' | 'hide\_fact' ) ('(open)')? (nameref + )
  1236     ;
  1237   \end{rail}
  1238 
  1239   Isabelle organizes any kind of name declarations (of types,
  1240   constants, theorems etc.) by separate hierarchically structured name
  1241   spaces.  Normally the user does not have to control the behavior of
  1242   name spaces by hand, yet the following commands provide some way to
  1243   do so.
  1244 
  1245   \begin{description}
  1246 
  1247   \item @{command "global"} and @{command "local"} change the current
  1248   name declaration mode.  Initially, theories start in @{command
  1249   "local"} mode, causing all names to be automatically qualified by
  1250   the theory name.  Changing this to @{command "global"} causes all
  1251   names to be declared without the theory prefix, until @{command
  1252   "local"} is declared again.
  1253   
  1254   Note that global names are prone to get hidden accidently later,
  1255   when qualified names of the same base name are introduced.
  1256   
  1257   \item @{command "hide_class"}~@{text names} fully removes class
  1258   declarations from a given name space; with the @{text "(open)"}
  1259   option, only the base name is hidden.  Global (unqualified) names
  1260   may never be hidden.
  1261 
  1262   Note that hiding name space accesses has no impact on logical
  1263   declarations --- they remain valid internally.  Entities that are no
  1264   longer accessible to the user are printed with the special qualifier
  1265   ``@{text "??"}'' prefixed to the full internal name.
  1266 
  1267   \item @{command "hide_type"}, @{command "hide_const"}, and @{command
  1268   "hide_fact"} are similar to @{command "hide_class"}, but hide types,
  1269   constants, and facts, respectively.
  1270   
  1271   \end{description}
  1272 *}
  1273 
  1274 end