doc-src/IsarRef/Thy/Spec.thy
author ballarin
Sun, 22 Nov 2009 15:37:14 +0100
changeset 33845 cff41e82e3df
parent 33516 0855a09bc5cf
child 33867 52643d0f856d
permissions -rw-r--r--
Updated locale documentation.
     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 "import + body"} prints the
   460   specified locale expression in a flattened form.  The notable
   461   special case @{command "print_locale"}~@{text loc} just prints the
   462   contents of the named locale, but keep in mind that type-inference
   463   will normalize type variables according to the usual alphabetical
   464   order.  The command omits @{element "notes"} elements by default.
   465   Use @{command "print_locale"}@{text "!"} to get them included.
   466 
   467   \item @{command "print_locales"} prints the names of all locales
   468   of the current theory.
   469 
   470   \item @{method intro_locales} and @{method unfold_locales}
   471   repeatedly expand all introduction rules of locale predicates of the
   472   theory.  While @{method intro_locales} only applies the @{text
   473   loc.intro} introduction rules and therefore does not decend to
   474   assumptions, @{method unfold_locales} is more aggressive and applies
   475   @{text loc_axioms.intro} as well.  Both methods are aware of locale
   476   specifications entailed by the context, both from target statements,
   477   and from interpretations (see below).  New goals that are entailed
   478   by the current context are discharged automatically.
   479 
   480   \end{description}
   481 *}
   482 
   483 
   484 subsection {* Locale interpretations *}
   485 
   486 text {*
   487   Locale expressions may be instantiated, and the instantiated facts
   488   added to the current context.  This requires a proof of the
   489   instantiated specification and is called \emph{locale
   490   interpretation}.  Interpretation is possible in locales @{command
   491   "sublocale"}, theories (command @{command "interpretation"}) and
   492   also within a proof body (command @{command "interpret"}).
   493 
   494   \begin{matharray}{rcl}
   495     @{command_def "sublocale"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   496     @{command_def "interpretation"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   497     @{command_def "interpret"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
   498     @{command_def "print_interps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   499   \end{matharray}
   500 
   501   \indexouternonterm{interp}
   502   \begin{rail}
   503     'sublocale' nameref ('<' | subseteq) localeexpr
   504     ;
   505     'interpretation' localeepxr equations?
   506     ;
   507     'interpret' localeexpr
   508     ;
   509     'print\_interps' nameref
   510     ;
   511     equations: 'where' (thmdecl? prop + 'and')
   512     ;
   513   \end{rail}
   514 
   515   \begin{description}
   516 
   517   \item @{command "sublocale"}~@{text "name \<subseteq> expr"}
   518   interprets @{text expr} in the locale @{text name}.  A proof that
   519   the specification of @{text name} implies the specification of
   520   @{text expr} is required.  As in the localized version of the
   521   theorem command, the proof is in the context of @{text name}.  After
   522   the proof obligation has been dischared, the facts of @{text expr}
   523   become part of locale @{text name} as \emph{derived} context
   524   elements and are available when the context @{text name} is
   525   subsequently entered.  Note that, like import, this is dynamic:
   526   facts added to a locale part of @{text expr} after interpretation
   527   become also available in @{text name}.
   528 
   529   Only specification fragments of @{text expr} that are not already
   530   part of @{text name} (be it imported, derived or a derived fragment
   531   of the import) are considered in this process.  This enables
   532   circular interpretations to the extent that no infinite chains are
   533   generated in the locale hierarchy.
   534 
   535   If interpretations of @{text name} exist in the current theory, the
   536   command adds interpretations for @{text expr} as well, with the same
   537   qualifier, although only for fragments of @{text expr} that are not
   538   interpreted in the theory already.
   539 
   540   \item @{command "interpretation"}~@{text "expr insts \<WHERE> eqns"}
   541   interprets @{text expr} in the theory.  The command generates proof
   542   obligations for the instantiated specifications (assumes and defines
   543   elements).  Once these are discharged by the user, instantiated
   544   facts are added to the theory in a post-processing phase.
   545 
   546   Additional equations, which are unfolded during
   547   post-processing, may be given after the keyword @{keyword "where"}.
   548   This is useful for interpreting concepts introduced through
   549   definition specification elements.  The equations must be proved.
   550 
   551   The command is aware of interpretations already active in the
   552   theory, but does not simplify the goal automatically.  In order to
   553   simplify the proof obligations use methods @{method intro_locales}
   554   or @{method unfold_locales}.  Post-processing is not applied to
   555   facts of interpretations that are already active.  This avoids
   556   duplication of interpreted facts, in particular.  Note that, in the
   557   case of a locale with import, parts of the interpretation may
   558   already be active.  The command will only process facts for new
   559   parts.
   560 
   561   Adding facts to locales has the effect of adding interpreted facts
   562   to the theory for all active interpretations also.  That is,
   563   interpretations dynamically participate in any facts added to
   564   locales.
   565 
   566   \item @{command "interpret"}~@{text "expr insts"}
   567   interprets @{text expr} in the proof context and is otherwise
   568   similar to interpretation in theories.
   569 
   570   \end{description}
   571 
   572   \begin{warn}
   573     Since attributes are applied to interpreted theorems,
   574     interpretation may modify the context of common proof tools, e.g.\
   575     the Simplifier or Classical Reasoner.  Since the behavior of such
   576     automated reasoning tools is \emph{not} stable under
   577     interpretation morphisms, manual declarations might have to be
   578     added to revert such declarations.
   579   \end{warn}
   580 
   581   \begin{warn}
   582     An interpretation in a theory may subsume previous
   583     interpretations.  This happens if the same specification fragment
   584     is interpreted twice and the instantiation of the second
   585     interpretation is more general than the interpretation of the
   586     first.  The locale package does not attempt to remove subsumed
   587     interpretations.
   588   \end{warn}
   589 *}
   590 
   591 
   592 section {* Classes \label{sec:class} *}
   593 
   594 text {*
   595   A class is a particular locale with \emph{exactly one} type variable
   596   @{text \<alpha>}.  Beyond the underlying locale, a corresponding type class
   597   is established which is interpreted logically as axiomatic type
   598   class \cite{Wenzel:1997:TPHOL} whose logical content are the
   599   assumptions of the locale.  Thus, classes provide the full
   600   generality of locales combined with the commodity of type classes
   601   (notably type-inference).  See \cite{isabelle-classes} for a short
   602   tutorial.
   603 
   604   \begin{matharray}{rcl}
   605     @{command_def "class"} & : & @{text "theory \<rightarrow> local_theory"} \\
   606     @{command_def "instantiation"} & : & @{text "theory \<rightarrow> local_theory"} \\
   607     @{command_def "instance"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   608     @{command_def "instance"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   609     @{command_def "subclass"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   610     @{command_def "print_classes"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   611     @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   612     @{method_def intro_classes} & : & @{text method} \\
   613   \end{matharray}
   614 
   615   \begin{rail}
   616     'class' name '=' ((superclassexpr '+' (contextelem+)) | superclassexpr | (contextelem+)) \\
   617       'begin'?
   618     ;
   619     'instantiation' (nameref + 'and') '::' arity 'begin'
   620     ;
   621     'instance'
   622     ;
   623     'instance' (nameref + 'and') '::' arity
   624     ;
   625     'subclass' target? nameref
   626     ;
   627     'instance' nameref ('<' | subseteq) nameref
   628     ;
   629     'print\_classes'
   630     ;
   631     'class\_deps'
   632     ;
   633 
   634     superclassexpr: nameref | (nameref '+' superclassexpr)
   635     ;
   636   \end{rail}
   637 
   638   \begin{description}
   639 
   640   \item @{command "class"}~@{text "c = superclasses + body"} defines
   641   a new class @{text c}, inheriting from @{text superclasses}.  This
   642   introduces a locale @{text c} with import of all locales @{text
   643   superclasses}.
   644 
   645   Any @{element "fixes"} in @{text body} are lifted to the global
   646   theory level (\emph{class operations} @{text "f\<^sub>1, \<dots>,
   647   f\<^sub>n"} of class @{text c}), mapping the local type parameter
   648   @{text \<alpha>} to a schematic type variable @{text "?\<alpha> :: c"}.
   649 
   650   Likewise, @{element "assumes"} in @{text body} are also lifted,
   651   mapping each local parameter @{text "f :: \<tau>[\<alpha>]"} to its
   652   corresponding global constant @{text "f :: \<tau>[?\<alpha> :: c]"}.  The
   653   corresponding introduction rule is provided as @{text
   654   c_class_axioms.intro}.  This rule should be rarely needed directly
   655   --- the @{method intro_classes} method takes care of the details of
   656   class membership proofs.
   657 
   658   \item @{command "instantiation"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s
   659   \<BEGIN>"} opens a theory target (cf.\ \secref{sec:target}) which
   660   allows to specify class operations @{text "f\<^sub>1, \<dots>, f\<^sub>n"} corresponding
   661   to sort @{text s} at the particular type instance @{text "(\<alpha>\<^sub>1 :: s\<^sub>1,
   662   \<dots>, \<alpha>\<^sub>n :: s\<^sub>n) t"}.  A plain @{command "instance"} command in the
   663   target body poses a goal stating these type arities.  The target is
   664   concluded by an @{command_ref (local) "end"} command.
   665 
   666   Note that a list of simultaneous type constructors may be given;
   667   this corresponds nicely to mutually recursive type definitions, e.g.\
   668   in Isabelle/HOL.
   669 
   670   \item @{command "instance"} in an instantiation target body sets
   671   up a goal stating the type arities claimed at the opening @{command
   672   "instantiation"}.  The proof would usually proceed by @{method
   673   intro_classes}, and then establish the characteristic theorems of
   674   the type classes involved.  After finishing the proof, the
   675   background theory will be augmented by the proven type arities.
   676 
   677   On the theory level, @{command "instance"}~@{text "t :: (s\<^sub>1, \<dots>,
   678   s\<^sub>n)s"} provides a convenient way to instantiate a type class with no
   679   need to specifify operations: one can continue with the
   680   instantiation proof immediately.
   681 
   682   \item @{command "subclass"}~@{text c} in a class context for class
   683   @{text d} sets up a goal stating that class @{text c} is logically
   684   contained in class @{text d}.  After finishing the proof, class
   685   @{text d} is proven to be subclass @{text c} and the locale @{text
   686   c} is interpreted into @{text d} simultaneously.
   687 
   688   A weakend form of this is available through a further variant of
   689   @{command instance}:  @{command instance}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} opens
   690   a proof that class @{text "c\<^isub>2"} implies @{text "c\<^isub>1"} without reference
   691   to the underlying locales;  this is useful if the properties to prove
   692   the logical connection are not sufficent on the locale level but on
   693   the theory level.
   694 
   695   \item @{command "print_classes"} prints all classes in the current
   696   theory.
   697 
   698   \item @{command "class_deps"} visualizes all classes and their
   699   subclass relations as a Hasse diagram.
   700 
   701   \item @{method intro_classes} repeatedly expands all class
   702   introduction rules of this theory.  Note that this method usually
   703   needs not be named explicitly, as it is already included in the
   704   default proof step (e.g.\ of @{command "proof"}).  In particular,
   705   instantiation of trivial (syntactic) classes may be performed by a
   706   single ``@{command ".."}'' proof step.
   707 
   708   \end{description}
   709 *}
   710 
   711 
   712 subsection {* The class target *}
   713 
   714 text {*
   715   %FIXME check
   716 
   717   A named context may refer to a locale (cf.\ \secref{sec:target}).
   718   If this locale is also a class @{text c}, apart from the common
   719   locale target behaviour the following happens.
   720 
   721   \begin{itemize}
   722 
   723   \item Local constant declarations @{text "g[\<alpha>]"} referring to the
   724   local type parameter @{text \<alpha>} and local parameters @{text "f[\<alpha>]"}
   725   are accompanied by theory-level constants @{text "g[?\<alpha> :: c]"}
   726   referring to theory-level class operations @{text "f[?\<alpha> :: c]"}.
   727 
   728   \item Local theorem bindings are lifted as are assumptions.
   729 
   730   \item Local syntax refers to local operations @{text "g[\<alpha>]"} and
   731   global operations @{text "g[?\<alpha> :: c]"} uniformly.  Type inference
   732   resolves ambiguities.  In rare cases, manual type annotations are
   733   needed.
   734   
   735   \end{itemize}
   736 *}
   737 
   738 
   739 subsection {* Old-style axiomatic type classes \label{sec:axclass} *}
   740 
   741 text {*
   742   \begin{matharray}{rcl}
   743     @{command_def "axclass"} & : & @{text "theory \<rightarrow> theory"}
   744   \end{matharray}
   745 
   746   Axiomatic type classes are Isabelle/Pure's primitive
   747   interface to type classes.  For practical
   748   applications, you should consider using classes
   749   (cf.~\secref{sec:classes}) which provide high level interface.
   750 
   751   \begin{rail}
   752     'axclass' classdecl (axmdecl prop +)
   753     ;
   754   \end{rail}
   755 
   756   \begin{description}
   757   
   758   \item @{command "axclass"}~@{text "c \<subseteq> c\<^sub>1, \<dots>, c\<^sub>n axms"} defines an
   759   axiomatic type class as the intersection of existing classes, with
   760   additional axioms holding.  Class axioms may not contain more than
   761   one type variable.  The class axioms (with implicit sort constraints
   762   added) are bound to the given names.  Furthermore a class
   763   introduction rule is generated (being bound as @{text
   764   c_class.intro}); this rule is employed by method @{method
   765   intro_classes} to support instantiation proofs of this class.
   766   
   767   The ``class axioms'' (which are derived from the internal class
   768   definition) are stored as theorems according to the given name
   769   specifications; the name space prefix @{text "c_class"} is added
   770   here.  The full collection of these facts is also stored as @{text
   771   c_class.axioms}.
   772   
   773   \end{description}
   774 *}
   775 
   776 
   777 section {* Unrestricted overloading *}
   778 
   779 text {*
   780   Isabelle/Pure's definitional schemes support certain forms of
   781   overloading (see \secref{sec:consts}).  Overloading means that a
   782   constant being declared as @{text "c :: \<alpha> decl"} may be
   783   defined separately on type instances
   784   @{text "c :: (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n) t decl"}
   785   for each type constructor @{text t}.  At most occassions
   786   overloading will be used in a Haskell-like fashion together with
   787   type classes by means of @{command "instantiation"} (see
   788   \secref{sec:class}).  Sometimes low-level overloading is desirable.
   789   The @{command "overloading"} target provides a convenient view for
   790   end-users.
   791 
   792   \begin{matharray}{rcl}
   793     @{command_def "overloading"} & : & @{text "theory \<rightarrow> local_theory"} \\
   794   \end{matharray}
   795 
   796   \begin{rail}
   797     'overloading' \\
   798     ( string ( '==' | equiv ) term ( '(' 'unchecked' ')' )? + ) 'begin'
   799   \end{rail}
   800 
   801   \begin{description}
   802 
   803   \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>"}
   804   opens a theory target (cf.\ \secref{sec:target}) which allows to
   805   specify constants with overloaded definitions.  These are identified
   806   by an explicitly given mapping from variable names @{text "x\<^sub>i"} to
   807   constants @{text "c\<^sub>i"} at particular type instances.  The
   808   definitions themselves are established using common specification
   809   tools, using the names @{text "x\<^sub>i"} as reference to the
   810   corresponding constants.  The target is concluded by @{command
   811   (local) "end"}.
   812 
   813   A @{text "(unchecked)"} option disables global dependency checks for
   814   the corresponding definition, which is occasionally useful for
   815   exotic overloading (see \secref{sec:consts} for a precise description).
   816   It is at the discretion of the user to avoid
   817   malformed theory specifications!
   818 
   819   \end{description}
   820 *}
   821 
   822 
   823 section {* Incorporating ML code \label{sec:ML} *}
   824 
   825 text {*
   826   \begin{matharray}{rcl}
   827     @{command_def "use"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   828     @{command_def "ML"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   829     @{command_def "ML_prf"} & : & @{text "proof \<rightarrow> proof"} \\
   830     @{command_def "ML_val"} & : & @{text "any \<rightarrow>"} \\
   831     @{command_def "ML_command"} & : & @{text "any \<rightarrow>"} \\
   832     @{command_def "setup"} & : & @{text "theory \<rightarrow> theory"} \\
   833     @{command_def "local_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   834     @{command_def "attribute_setup"} & : & @{text "theory \<rightarrow> theory"} \\
   835   \end{matharray}
   836 
   837   \begin{mldecls}
   838     @{index_ML bind_thms: "string * thm list -> unit"} \\
   839     @{index_ML bind_thm: "string * thm -> unit"} \\
   840   \end{mldecls}
   841 
   842   \begin{rail}
   843     'use' name
   844     ;
   845     ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup' | 'local\_setup') text
   846     ;
   847     'attribute\_setup' name '=' text text
   848     ;
   849   \end{rail}
   850 
   851   \begin{description}
   852 
   853   \item @{command "use"}~@{text "file"} reads and executes ML
   854   commands from @{text "file"}.  The current theory context is passed
   855   down to the ML toplevel and may be modified, using @{ML
   856   "Context.>>"} or derived ML commands.  The file name is checked with
   857   the @{keyword_ref "uses"} dependency declaration given in the theory
   858   header (see also \secref{sec:begin-thy}).
   859 
   860   Top-level ML bindings are stored within the (global or local) theory
   861   context.
   862   
   863   \item @{command "ML"}~@{text "text"} is similar to @{command "use"},
   864   but executes ML commands directly from the given @{text "text"}.
   865   Top-level ML bindings are stored within the (global or local) theory
   866   context.
   867 
   868   \item @{command "ML_prf"} is analogous to @{command "ML"} but works
   869   within a proof context.
   870 
   871   Top-level ML bindings are stored within the proof context in a
   872   purely sequential fashion, disregarding the nested proof structure.
   873   ML bindings introduced by @{command "ML_prf"} are discarded at the
   874   end of the proof.
   875 
   876   \item @{command "ML_val"} and @{command "ML_command"} are diagnostic
   877   versions of @{command "ML"}, which means that the context may not be
   878   updated.  @{command "ML_val"} echos the bindings produced at the ML
   879   toplevel, but @{command "ML_command"} is silent.
   880   
   881   \item @{command "setup"}~@{text "text"} changes the current theory
   882   context by applying @{text "text"}, which refers to an ML expression
   883   of type @{ML_type "theory -> theory"}.  This enables to initialize
   884   any object-logic specific tools and packages written in ML, for
   885   example.
   886 
   887   \item @{command "local_setup"} is similar to @{command "setup"} for
   888   a local theory context, and an ML expression of type @{ML_type
   889   "local_theory -> local_theory"}.  This allows to
   890   invoke local theory specification packages without going through
   891   concrete outer syntax, for example.
   892 
   893   \item @{command "attribute_setup"}~@{text "name = text description"}
   894   defines an attribute in the current theory.  The given @{text
   895   "text"} has to be an ML expression of type
   896   @{ML_type "attribute context_parser"}, cf.\ basic parsers defined in
   897   structure @{ML_struct Args} and @{ML_struct Attrib}.
   898 
   899   In principle, attributes can operate both on a given theorem and the
   900   implicit context, although in practice only one is modified and the
   901   other serves as parameter.  Here are examples for these two cases:
   902 
   903   \end{description}
   904 *}
   905 
   906     attribute_setup my_rule = {*
   907       Attrib.thms >> (fn ths =>
   908         Thm.rule_attribute (fn context: Context.generic => fn th: thm =>
   909           let val th' = th OF ths
   910           in th' end)) *}  "my rule"
   911 
   912     attribute_setup my_declaration = {*
   913       Attrib.thms >> (fn ths =>
   914         Thm.declaration_attribute (fn th: thm => fn context: Context.generic =>
   915           let val context' = context
   916           in context' end)) *}  "my declaration"
   917 
   918 text {*
   919   \begin{description}
   920 
   921   \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
   922   theorems produced in ML both in the theory context and the ML
   923   toplevel, associating it with the provided name.  Theorems are put
   924   into a global ``standard'' format before being stored.
   925 
   926   \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a
   927   singleton theorem.
   928   
   929   \end{description}
   930 *}
   931 
   932 
   933 section {* Primitive specification elements *}
   934 
   935 subsection {* Type classes and sorts \label{sec:classes} *}
   936 
   937 text {*
   938   \begin{matharray}{rcll}
   939     @{command_def "classes"} & : & @{text "theory \<rightarrow> theory"} \\
   940     @{command_def "classrel"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   941     @{command_def "defaultsort"} & : & @{text "theory \<rightarrow> theory"} \\
   942     @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   943   \end{matharray}
   944 
   945   \begin{rail}
   946     'classes' (classdecl +)
   947     ;
   948     'classrel' (nameref ('<' | subseteq) nameref + 'and')
   949     ;
   950     'defaultsort' sort
   951     ;
   952   \end{rail}
   953 
   954   \begin{description}
   955 
   956   \item @{command "classes"}~@{text "c \<subseteq> c\<^sub>1, \<dots>, c\<^sub>n"} declares class
   957   @{text c} to be a subclass of existing classes @{text "c\<^sub>1, \<dots>, c\<^sub>n"}.
   958   Isabelle implicitly maintains the transitive closure of the class
   959   hierarchy.  Cyclic class structures are not permitted.
   960 
   961   \item @{command "classrel"}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} states subclass
   962   relations between existing classes @{text "c\<^sub>1"} and @{text "c\<^sub>2"}.
   963   This is done axiomatically!  The @{command_ref "instance"} command
   964   (see \secref{sec:axclass}) provides a way to introduce proven class
   965   relations.
   966 
   967   \item @{command "defaultsort"}~@{text s} makes sort @{text s} the
   968   new default sort for any type variable that is given explicitly in
   969   the text, but lacks a sort constraint (wrt.\ the current context).
   970   Type variables generated by type inference are not affected.
   971 
   972   Usually the default sort is only changed when defining a new
   973   object-logic.  For example, the default sort in Isabelle/HOL is
   974   @{text type}, the class of all HOL types.  %FIXME sort antiq?
   975 
   976   When merging theories, the default sorts of the parents are
   977   logically intersected, i.e.\ the representations as lists of classes
   978   are joined.
   979 
   980   \item @{command "class_deps"} visualizes the subclass relation,
   981   using Isabelle's graph browser tool (see also \cite{isabelle-sys}).
   982 
   983   \end{description}
   984 *}
   985 
   986 
   987 subsection {* Types and type abbreviations \label{sec:types-pure} *}
   988 
   989 text {*
   990   \begin{matharray}{rcll}
   991     @{command_def "types"} & : & @{text "theory \<rightarrow> theory"} \\
   992     @{command_def "typedecl"} & : & @{text "theory \<rightarrow> theory"} \\
   993     @{command_def "arities"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   994   \end{matharray}
   995 
   996   \begin{rail}
   997     'types' (typespec '=' type infix? +)
   998     ;
   999     'typedecl' typespec infix?
  1000     ;
  1001     'arities' (nameref '::' arity +)
  1002     ;
  1003   \end{rail}
  1004 
  1005   \begin{description}
  1006 
  1007   \item @{command "types"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = \<tau>"} introduces a
  1008   \emph{type synonym} @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} for the existing type
  1009   @{text "\<tau>"}.  Unlike actual type definitions, as are available in
  1010   Isabelle/HOL for example, type synonyms are merely syntactic
  1011   abbreviations without any logical significance.  Internally, type
  1012   synonyms are fully expanded.
  1013   
  1014   \item @{command "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} declares a new
  1015   type constructor @{text t}.  If the object-logic defines a base sort
  1016   @{text s}, then the constructor is declared to operate on that, via
  1017   the axiomatic specification @{command arities}~@{text "t :: (s, \<dots>,
  1018   s)s"}.
  1019 
  1020   \item @{command "arities"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s"} augments
  1021   Isabelle's order-sorted signature of types by new type constructor
  1022   arities.  This is done axiomatically!  The @{command_ref "instance"}
  1023   command (see \secref{sec:axclass}) provides a way to introduce
  1024   proven type arities.
  1025 
  1026   \end{description}
  1027 *}
  1028 
  1029 
  1030 subsection {* Co-regularity of type classes and arities *}
  1031 
  1032 text {* The class relation together with the collection of
  1033   type-constructor arities must obey the principle of
  1034   \emph{co-regularity} as defined below.
  1035 
  1036   \medskip For the subsequent formulation of co-regularity we assume
  1037   that the class relation is closed by transitivity and reflexivity.
  1038   Moreover the collection of arities @{text "t :: (\<^vec>s)c"} is
  1039   completed such that @{text "t :: (\<^vec>s)c"} and @{text "c \<subseteq> c'"}
  1040   implies @{text "t :: (\<^vec>s)c'"} for all such declarations.
  1041 
  1042   Treating sorts as finite sets of classes (meaning the intersection),
  1043   the class relation @{text "c\<^sub>1 \<subseteq> c\<^sub>2"} is extended to sorts as
  1044   follows:
  1045   \[
  1046     @{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"}
  1047   \]
  1048 
  1049   This relation on sorts is further extended to tuples of sorts (of
  1050   the same length) in the component-wise way.
  1051 
  1052   \smallskip Co-regularity of the class relation together with the
  1053   arities relation means:
  1054   \[
  1055     @{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"}
  1056   \]
  1057   \noindent for all such arities.  In other words, whenever the result
  1058   classes of some type-constructor arities are related, then the
  1059   argument sorts need to be related in the same way.
  1060 
  1061   \medskip Co-regularity is a very fundamental property of the
  1062   order-sorted algebra of types.  For example, it entails principle
  1063   types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}.
  1064 *}
  1065 
  1066 
  1067 subsection {* Constants and definitions \label{sec:consts} *}
  1068 
  1069 text {*
  1070   Definitions essentially express abbreviations within the logic.  The
  1071   simplest form of a definition is @{text "c :: \<sigma> \<equiv> t"}, where @{text
  1072   c} is a newly declared constant.  Isabelle also allows derived forms
  1073   where the arguments of @{text c} appear on the left, abbreviating a
  1074   prefix of @{text \<lambda>}-abstractions, e.g.\ @{text "c \<equiv> \<lambda>x y. t"} may be
  1075   written more conveniently as @{text "c x y \<equiv> t"}.  Moreover,
  1076   definitions may be weakened by adding arbitrary pre-conditions:
  1077   @{text "A \<Longrightarrow> c x y \<equiv> t"}.
  1078 
  1079   \medskip The built-in well-formedness conditions for definitional
  1080   specifications are:
  1081 
  1082   \begin{itemize}
  1083 
  1084   \item Arguments (on the left-hand side) must be distinct variables.
  1085 
  1086   \item All variables on the right-hand side must also appear on the
  1087   left-hand side.
  1088 
  1089   \item All type variables on the right-hand side must also appear on
  1090   the left-hand side; this prohibits @{text "0 :: nat \<equiv> length ([] ::
  1091   \<alpha> list)"} for example.
  1092 
  1093   \item The definition must not be recursive.  Most object-logics
  1094   provide definitional principles that can be used to express
  1095   recursion safely.
  1096 
  1097   \end{itemize}
  1098 
  1099   The right-hand side of overloaded definitions may mention overloaded constants
  1100   recursively at type instances corresponding to the immediate
  1101   argument types @{text "\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n"}.  Incomplete
  1102   specification patterns impose global constraints on all occurrences,
  1103   e.g.\ @{text "d :: \<alpha> \<times> \<alpha>"} on the left-hand side means that all
  1104   corresponding occurrences on some right-hand side need to be an
  1105   instance of this, general @{text "d :: \<alpha> \<times> \<beta>"} will be disallowed.
  1106 
  1107   \begin{matharray}{rcl}
  1108     @{command_def "consts"} & : & @{text "theory \<rightarrow> theory"} \\
  1109     @{command_def "defs"} & : & @{text "theory \<rightarrow> theory"} \\
  1110     @{command_def "constdefs"} & : & @{text "theory \<rightarrow> theory"} \\
  1111   \end{matharray}
  1112 
  1113   \begin{rail}
  1114     'consts' ((name '::' type mixfix?) +)
  1115     ;
  1116     'defs' ('(' 'unchecked'? 'overloaded'? ')')? \\ (axmdecl prop +)
  1117     ;
  1118   \end{rail}
  1119 
  1120   \begin{rail}
  1121     'constdefs' structs? (constdecl? constdef +)
  1122     ;
  1123 
  1124     structs: '(' 'structure' (vars + 'and') ')'
  1125     ;
  1126     constdecl:  ((name '::' type mixfix | name '::' type | name mixfix) 'where'?) | name 'where'
  1127     ;
  1128     constdef: thmdecl? prop
  1129     ;
  1130   \end{rail}
  1131 
  1132   \begin{description}
  1133 
  1134   \item @{command "consts"}~@{text "c :: \<sigma>"} declares constant @{text
  1135   c} to have any instance of type scheme @{text \<sigma>}.  The optional
  1136   mixfix annotations may attach concrete syntax to the constants
  1137   declared.
  1138   
  1139   \item @{command "defs"}~@{text "name: eqn"} introduces @{text eqn}
  1140   as a definitional axiom for some existing constant.
  1141   
  1142   The @{text "(unchecked)"} option disables global dependency checks
  1143   for this definition, which is occasionally useful for exotic
  1144   overloading.  It is at the discretion of the user to avoid malformed
  1145   theory specifications!
  1146   
  1147   The @{text "(overloaded)"} option declares definitions to be
  1148   potentially overloaded.  Unless this option is given, a warning
  1149   message would be issued for any definitional equation with a more
  1150   special type than that of the corresponding constant declaration.
  1151   
  1152   \item @{command "constdefs"} combines constant declarations and
  1153   definitions, with type-inference taking care of the most general
  1154   typing of the given specification (the optional type constraint may
  1155   refer to type-inference dummies ``@{text _}'' as usual).  The
  1156   resulting type declaration needs to agree with that of the
  1157   specification; overloading is \emph{not} supported here!
  1158   
  1159   The constant name may be omitted altogether, if neither type nor
  1160   syntax declarations are given.  The canonical name of the
  1161   definitional axiom for constant @{text c} will be @{text c_def},
  1162   unless specified otherwise.  Also note that the given list of
  1163   specifications is processed in a strictly sequential manner, with
  1164   type-checking being performed independently.
  1165   
  1166   An optional initial context of @{text "(structure)"} declarations
  1167   admits use of indexed syntax, using the special symbol @{verbatim
  1168   "\<index>"} (printed as ``@{text "\<index>"}'').  The latter concept is
  1169   particularly useful with locales (see also \secref{sec:locale}).
  1170 
  1171   \end{description}
  1172 *}
  1173 
  1174 
  1175 section {* Axioms and theorems \label{sec:axms-thms} *}
  1176 
  1177 text {*
  1178   \begin{matharray}{rcll}
  1179     @{command_def "axioms"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
  1180     @{command_def "lemmas"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1181     @{command_def "theorems"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1182   \end{matharray}
  1183 
  1184   \begin{rail}
  1185     'axioms' (axmdecl prop +)
  1186     ;
  1187     ('lemmas' | 'theorems') target? (thmdef? thmrefs + 'and')
  1188     ;
  1189   \end{rail}
  1190 
  1191   \begin{description}
  1192   
  1193   \item @{command "axioms"}~@{text "a: \<phi>"} introduces arbitrary
  1194   statements as axioms of the meta-logic.  In fact, axioms are
  1195   ``axiomatic theorems'', and may be referred later just as any other
  1196   theorem.
  1197   
  1198   Axioms are usually only introduced when declaring new logical
  1199   systems.  Everyday work is typically done the hard way, with proper
  1200   definitions and proven theorems.
  1201   
  1202   \item @{command "lemmas"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"} retrieves and stores
  1203   existing facts in the theory context, or the specified target
  1204   context (see also \secref{sec:target}).  Typical applications would
  1205   also involve attributes, to declare Simplifier rules, for example.
  1206   
  1207   \item @{command "theorems"} is essentially the same as @{command
  1208   "lemmas"}, but marks the result as a different kind of facts.
  1209 
  1210   \end{description}
  1211 *}
  1212 
  1213 
  1214 section {* Oracles *}
  1215 
  1216 text {* Oracles allow Isabelle to take advantage of external reasoners
  1217   such as arithmetic decision procedures, model checkers, fast
  1218   tautology checkers or computer algebra systems.  Invoked as an
  1219   oracle, an external reasoner can create arbitrary Isabelle theorems.
  1220 
  1221   It is the responsibility of the user to ensure that the external
  1222   reasoner is as trustworthy as the application requires.  Another
  1223   typical source of errors is the linkup between Isabelle and the
  1224   external tool, not just its concrete implementation, but also the
  1225   required translation between two different logical environments.
  1226 
  1227   Isabelle merely guarantees well-formedness of the propositions being
  1228   asserted, and records within the internal derivation object how
  1229   presumed theorems depend on unproven suppositions.
  1230 
  1231   \begin{matharray}{rcl}
  1232     @{command_def "oracle"} & : & @{text "theory \<rightarrow> theory"} \\
  1233   \end{matharray}
  1234 
  1235   \begin{rail}
  1236     'oracle' name '=' text
  1237     ;
  1238   \end{rail}
  1239 
  1240   \begin{description}
  1241 
  1242   \item @{command "oracle"}~@{text "name = text"} turns the given ML
  1243   expression @{text "text"} of type @{ML_text "'a -> cterm"} into an
  1244   ML function of type @{ML_text "'a -> thm"}, which is bound to the
  1245   global identifier @{ML_text name}.  This acts like an infinitary
  1246   specification of axioms!  Invoking the oracle only works within the
  1247   scope of the resulting theory.
  1248 
  1249   \end{description}
  1250 
  1251   See @{"file" "~~/src/FOL/ex/Iff_Oracle.thy"} for a worked example of
  1252   defining a new primitive rule as oracle, and turning it into a proof
  1253   method.
  1254 *}
  1255 
  1256 
  1257 section {* Name spaces *}
  1258 
  1259 text {*
  1260   \begin{matharray}{rcl}
  1261     @{command_def "global"} & : & @{text "theory \<rightarrow> theory"} \\
  1262     @{command_def "local"} & : & @{text "theory \<rightarrow> theory"} \\
  1263     @{command_def "hide"} & : & @{text "theory \<rightarrow> theory"} \\
  1264   \end{matharray}
  1265 
  1266   \begin{rail}
  1267     'hide' ('(open)')? name (nameref + )
  1268     ;
  1269   \end{rail}
  1270 
  1271   Isabelle organizes any kind of name declarations (of types,
  1272   constants, theorems etc.) by separate hierarchically structured name
  1273   spaces.  Normally the user does not have to control the behavior of
  1274   name spaces by hand, yet the following commands provide some way to
  1275   do so.
  1276 
  1277   \begin{description}
  1278 
  1279   \item @{command "global"} and @{command "local"} change the current
  1280   name declaration mode.  Initially, theories start in @{command
  1281   "local"} mode, causing all names to be automatically qualified by
  1282   the theory name.  Changing this to @{command "global"} causes all
  1283   names to be declared without the theory prefix, until @{command
  1284   "local"} is declared again.
  1285   
  1286   Note that global names are prone to get hidden accidently later,
  1287   when qualified names of the same base name are introduced.
  1288   
  1289   \item @{command "hide"}~@{text "space names"} fully removes
  1290   declarations from a given name space (which may be @{text "class"},
  1291   @{text "type"}, @{text "const"}, or @{text "fact"}); with the @{text
  1292   "(open)"} option, only the base name is hidden.  Global
  1293   (unqualified) names may never be hidden.
  1294   
  1295   Note that hiding name space accesses has no impact on logical
  1296   declarations --- they remain valid internally.  Entities that are no
  1297   longer accessible to the user are printed with the special qualifier
  1298   ``@{text "??"}'' prefixed to the full internal name.
  1299 
  1300   \end{description}
  1301 *}
  1302 
  1303 end