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