doc-src/IsarRef/Thy/Spec.thy
author wenzelm
Sat, 17 Mar 2012 22:46:19 +0100
changeset 47870 1c3c185bab4e
parent 46471 1bbbac9a0cb0
child 47985 7c9e31ffcd9e
permissions -rw-r--r--
more precise syntax;
     1 theory Spec
     2 imports Base Main
     3 begin
     4 
     5 chapter {* 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   @{rail "
    54     @@{command theory} @{syntax name} \\ @'imports' (@{syntax name} +) uses? @'begin'
    55     ;
    56 
    57     uses: @'uses' ((@{syntax name} | @{syntax parname}) +)
    58   "}
    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   @{rail "
   106     @@{command context} @{syntax nameref} @'begin'
   107     ;
   108 
   109     @{syntax_def target}: '(' @'in' @{syntax nameref} ')'
   110   "}
   111 
   112   \begin{description}
   113   
   114   \item @{command "context"}~@{text "c \<BEGIN>"} recommences an
   115   existing locale or class context @{text c}.  Note that locale and
   116   class definitions allow to include the @{keyword "begin"} keyword as
   117   well, in order to continue the local theory immediately after the
   118   initial specification.
   119   
   120   \item @{command (local) "end"} concludes the current local theory
   121   and continues the enclosing global theory.  Note that a global
   122   @{command (global) "end"} has a different meaning: it concludes the
   123   theory itself (\secref{sec:begin-thy}).
   124   
   125   \item @{text "("}@{keyword_def "in"}~@{text "c)"} given after any
   126   local theory command specifies an immediate target, e.g.\
   127   ``@{command "definition"}~@{text "(\<IN> c) \<dots>"}'' or ``@{command
   128   "theorem"}~@{text "(\<IN> c) \<dots>"}''.  This works both in a local or
   129   global theory context; the current target context will be suspended
   130   for this command only.  Note that ``@{text "(\<IN> -)"}'' will
   131   always produce a global result independently of the current target
   132   context.
   133 
   134   \end{description}
   135 
   136   The exact meaning of results produced within a local theory context
   137   depends on the underlying target infrastructure (locale, type class
   138   etc.).  The general idea is as follows, considering a context named
   139   @{text c} with parameter @{text x} and assumption @{text "A[x]"}.
   140   
   141   Definitions are exported by introducing a global version with
   142   additional arguments; a syntactic abbreviation links the long form
   143   with the abstract version of the target context.  For example,
   144   @{text "a \<equiv> t[x]"} becomes @{text "c.a ?x \<equiv> t[?x]"} at the theory
   145   level (for arbitrary @{text "?x"}), together with a local
   146   abbreviation @{text "c \<equiv> c.a x"} in the target context (for the
   147   fixed parameter @{text x}).
   148 
   149   Theorems are exported by discharging the assumptions and
   150   generalizing the parameters of the context.  For example, @{text "a:
   151   B[x]"} becomes @{text "c.a: A[?x] \<Longrightarrow> B[?x]"}, again for arbitrary
   152   @{text "?x"}.
   153 *}
   154 
   155 
   156 section {* Basic specification elements *}
   157 
   158 text {*
   159   \begin{matharray}{rcll}
   160     @{command_def "axiomatization"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   161     @{command_def "definition"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   162     @{attribute_def "defn"} & : & @{text attribute} \\
   163     @{command_def "abbreviation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   164     @{command_def "print_abbrevs"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow> "} \\
   165   \end{matharray}
   166 
   167   These specification mechanisms provide a slightly more abstract view
   168   than the underlying primitives of @{command "consts"}, @{command
   169   "defs"} (see \secref{sec:consts}), and @{command "axioms"} (see
   170   \secref{sec:axms-thms}).  In particular, type-inference is commonly
   171   available, and result names need not be given.
   172 
   173   @{rail "
   174     @@{command axiomatization} @{syntax \"fixes\"}? (@'where' specs)?
   175     ;
   176     @@{command definition} @{syntax target}? \\
   177       (decl @'where')? @{syntax thmdecl}? @{syntax prop}
   178     ;
   179     @@{command abbreviation} @{syntax target}? @{syntax mode}? \\
   180       (decl @'where')? @{syntax prop}
   181     ;
   182 
   183     @{syntax_def \"fixes\"}: ((@{syntax name} ('::' @{syntax type})?
   184       @{syntax mixfix}? | @{syntax vars}) + @'and')
   185     ;
   186     specs: (@{syntax thmdecl}? @{syntax props} + @'and')
   187     ;
   188     decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}?
   189   "}
   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 "syntax_declaration"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   257     @{command_def "declare"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   258   \end{matharray}
   259 
   260   @{rail "
   261     (@@{command declaration} | @@{command syntax_declaration})
   262       ('(' @'pervasive' ')')? \\ @{syntax target}? @{syntax text}
   263     ;
   264     @@{command declare} @{syntax target}? (@{syntax thmrefs} + @'and')
   265   "}
   266 
   267   \begin{description}
   268 
   269   \item @{command "declaration"}~@{text d} adds the declaration
   270   function @{text d} of ML type @{ML_type declaration}, to the current
   271   local theory under construction.  In later application contexts, the
   272   function is transformed according to the morphisms being involved in
   273   the interpretation hierarchy.
   274 
   275   If the @{text "(pervasive)"} option is given, the corresponding
   276   declaration is applied to all possible contexts involved, including
   277   the global background theory.
   278 
   279   \item @{command "syntax_declaration"} is similar to @{command
   280   "declaration"}, but is meant to affect only ``syntactic'' tools by
   281   convention (such as notation and type-checking information).
   282 
   283   \item @{command "declare"}~@{text thms} declares theorems to the
   284   current local theory context.  No theorem binding is involved here,
   285   unlike @{command "theorems"} or @{command "lemmas"} (cf.\
   286   \secref{sec:axms-thms}), so @{command "declare"} only has the effect
   287   of applying attributes as included in the theorem specification.
   288 
   289   \end{description}
   290 *}
   291 
   292 
   293 section {* Locales \label{sec:locale} *}
   294 
   295 text {*
   296   Locales are parametric named local contexts, consisting of a list of
   297   declaration elements that are modeled after the Isar proof context
   298   commands (cf.\ \secref{sec:proof-context}).
   299 *}
   300 
   301 
   302 subsection {* Locale expressions \label{sec:locale-expr} *}
   303 
   304 text {*
   305   A \emph{locale expression} denotes a structured context composed of
   306   instances of existing locales.  The context consists of a list of
   307   instances of declaration elements from the locales.  Two locale
   308   instances are equal if they are of the same locale and the
   309   parameters are instantiated with equivalent terms.  Declaration
   310   elements from equal instances are never repeated, thus avoiding
   311   duplicate declarations.
   312 
   313   @{rail "
   314     @{syntax_def locale_expr}: (instance + '+') (@'for' (@{syntax \"fixes\"} + @'and'))?
   315     ;
   316     instance: (qualifier ':')? @{syntax nameref} (pos_insts | named_insts)
   317     ;
   318     qualifier: @{syntax name} ('?' | '!')?
   319     ;
   320     pos_insts: ('_' | @{syntax term})*
   321     ;
   322     named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and')
   323   "}
   324 
   325   A locale instance consists of a reference to a locale and either
   326   positional or named parameter instantiations.  Identical
   327   instantiations (that is, those that instante a parameter by itself)
   328   may be omitted.  The notation `@{text "_"}' enables to omit the
   329   instantiation for a parameter inside a positional instantiation.
   330 
   331   Terms in instantiations are from the context the locale expressions
   332   is declared in.  Local names may be added to this context with the
   333   optional for clause.  In addition, syntax declarations from one
   334   instance are effective when parsing subsequent instances of the same
   335   expression.
   336 
   337   Instances have an optional qualifier which applies to names in
   338   declarations.  Names include local definitions and theorem names.
   339   If present, the qualifier itself is either optional
   340   (``\texttt{?}''), which means that it may be omitted on input of the
   341   qualified name, or mandatory (``\texttt{!}'').  If neither
   342   ``\texttt{?}'' nor ``\texttt{!}'' are present, the command's default
   343   is used.  For @{command "interpretation"} and @{command "interpret"}
   344   the default is ``mandatory'', for @{command "locale"} and @{command
   345   "sublocale"} the default is ``optional''.
   346 *}
   347 
   348 
   349 subsection {* Locale declarations *}
   350 
   351 text {*
   352   \begin{matharray}{rcl}
   353     @{command_def "locale"} & : & @{text "theory \<rightarrow> local_theory"} \\
   354     @{command_def "print_locale"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   355     @{command_def "print_locales"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   356     @{method_def intro_locales} & : & @{text method} \\
   357     @{method_def unfold_locales} & : & @{text method} \\
   358   \end{matharray}
   359 
   360   \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes}
   361   \indexisarelem{defines}\indexisarelem{notes}
   362   @{rail "
   363     @@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'?
   364     ;
   365     @@{command print_locale} '!'? @{syntax nameref}
   366     ;
   367     @{syntax_def locale}: @{syntax context_elem}+ |
   368       @{syntax locale_expr} ('+' (@{syntax context_elem}+))?
   369     ;
   370     @{syntax_def context_elem}:
   371       @'fixes' (@{syntax \"fixes\"} + @'and') |
   372       @'constrains' (@{syntax name} '::' @{syntax type} + @'and') |
   373       @'assumes' (@{syntax props} + @'and') |
   374       @'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') |
   375       @'notes' (@{syntax thmdef}? @{syntax thmrefs} + @'and')
   376   "}
   377 
   378   \begin{description}
   379   
   380   \item @{command "locale"}~@{text "loc = import + body"} defines a
   381   new locale @{text loc} as a context consisting of a certain view of
   382   existing locales (@{text import}) plus some additional elements
   383   (@{text body}).  Both @{text import} and @{text body} are optional;
   384   the degenerate form @{command "locale"}~@{text loc} defines an empty
   385   locale, which may still be useful to collect declarations of facts
   386   later on.  Type-inference on locale expressions automatically takes
   387   care of the most general typing that the combined context elements
   388   may acquire.
   389 
   390   The @{text import} consists of a structured locale expression; see
   391   \secref{sec:proof-context} above.  Its for clause defines the local
   392   parameters of the @{text import}.  In addition, locale parameters
   393   whose instantance is omitted automatically extend the (possibly
   394   empty) for clause: they are inserted at its beginning.  This means
   395   that these parameters may be referred to from within the expression
   396   and also in the subsequent context elements and provides a
   397   notational convenience for the inheritance of parameters in locale
   398   declarations.
   399 
   400   The @{text body} consists of context elements.
   401 
   402   \begin{description}
   403 
   404   \item @{element "fixes"}~@{text "x :: \<tau> (mx)"} declares a local
   405   parameter of type @{text \<tau>} and mixfix annotation @{text mx} (both
   406   are optional).  The special syntax declaration ``@{text
   407   "(\<STRUCTURE>)"}'' means that @{text x} may be referenced
   408   implicitly in this context.
   409 
   410   \item @{element "constrains"}~@{text "x :: \<tau>"} introduces a type
   411   constraint @{text \<tau>} on the local parameter @{text x}.  This
   412   element is deprecated.  The type constraint should be introduced in
   413   the for clause or the relevant @{element "fixes"} element.
   414 
   415   \item @{element "assumes"}~@{text "a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"}
   416   introduces local premises, similar to @{command "assume"} within a
   417   proof (cf.\ \secref{sec:proof-context}).
   418 
   419   \item @{element "defines"}~@{text "a: x \<equiv> t"} defines a previously
   420   declared parameter.  This is similar to @{command "def"} within a
   421   proof (cf.\ \secref{sec:proof-context}), but @{element "defines"}
   422   takes an equational proposition instead of variable-term pair.  The
   423   left-hand side of the equation may have additional arguments, e.g.\
   424   ``@{element "defines"}~@{text "f x\<^sub>1 \<dots> x\<^sub>n \<equiv> t"}''.
   425 
   426   \item @{element "notes"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"}
   427   reconsiders facts within a local context.  Most notably, this may
   428   include arbitrary declarations in any attribute specifications
   429   included here, e.g.\ a local @{attribute simp} rule.
   430 
   431   The initial @{text import} specification of a locale expression
   432   maintains a dynamic relation to the locales being referenced
   433   (benefiting from any later fact declarations in the obvious manner).
   434 
   435   \end{description}
   436   
   437   Note that ``@{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"}'' patterns given
   438   in the syntax of @{element "assumes"} and @{element "defines"} above
   439   are illegal in locale definitions.  In the long goal format of
   440   \secref{sec:goals}, term bindings may be included as expected,
   441   though.
   442   
   443   \medskip Locale specifications are ``closed up'' by
   444   turning the given text into a predicate definition @{text
   445   loc_axioms} and deriving the original assumptions as local lemmas
   446   (modulo local definitions).  The predicate statement covers only the
   447   newly specified assumptions, omitting the content of included locale
   448   expressions.  The full cumulative view is only provided on export,
   449   involving another predicate @{text loc} that refers to the complete
   450   specification text.
   451   
   452   In any case, the predicate arguments are those locale parameters
   453   that actually occur in the respective piece of text.  Also note that
   454   these predicates operate at the meta-level in theory, but the locale
   455   packages attempts to internalize statements according to the
   456   object-logic setup (e.g.\ replacing @{text \<And>} by @{text \<forall>}, and
   457   @{text "\<Longrightarrow>"} by @{text "\<longrightarrow>"} in HOL; see also
   458   \secref{sec:object-logic}).  Separate introduction rules @{text
   459   loc_axioms.intro} and @{text loc.intro} are provided as well.
   460   
   461   \item @{command "print_locale"}~@{text "locale"} prints the
   462   contents of the named locale.  The command omits @{element "notes"}
   463   elements by default.  Use @{command "print_locale"}@{text "!"} to
   464   have them included.
   465 
   466   \item @{command "print_locales"} prints the names of all locales
   467   of the current theory.
   468 
   469   \item @{method intro_locales} and @{method unfold_locales}
   470   repeatedly expand all introduction rules of locale predicates of the
   471   theory.  While @{method intro_locales} only applies the @{text
   472   loc.intro} introduction rules and therefore does not decend to
   473   assumptions, @{method unfold_locales} is more aggressive and applies
   474   @{text loc_axioms.intro} as well.  Both methods are aware of locale
   475   specifications entailed by the context, both from target statements,
   476   and from interpretations (see below).  New goals that are entailed
   477   by the current context are discharged automatically.
   478 
   479   \end{description}
   480 *}
   481 
   482 
   483 subsection {* Locale interpretations *}
   484 
   485 text {*
   486   Locale expressions may be instantiated, and the instantiated facts
   487   added to the current context.  This requires a proof of the
   488   instantiated specification and is called \emph{locale
   489   interpretation}.  Interpretation is possible in locales @{command
   490   "sublocale"}, theories (command @{command "interpretation"}) and
   491   also within a proof body (command @{command "interpret"}).
   492 
   493   \begin{matharray}{rcl}
   494     @{command_def "interpretation"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   495     @{command_def "interpret"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
   496     @{command_def "sublocale"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   497     @{command_def "print_dependencies"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   498     @{command_def "print_interps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   499   \end{matharray}
   500 
   501   @{rail "
   502     @@{command interpretation} @{syntax locale_expr} equations?
   503     ;
   504     @@{command interpret} @{syntax locale_expr} equations?
   505     ;
   506     @@{command sublocale} @{syntax nameref} ('<' | '\<subseteq>') @{syntax locale_expr} \\
   507       equations?
   508     ;
   509     @@{command print_dependencies} '!'? @{syntax locale_expr}
   510     ;
   511     @@{command print_interps} @{syntax nameref}
   512     ;
   513 
   514     equations: @'where' (@{syntax thmdecl}? @{syntax prop} + @'and')
   515   "}
   516 
   517   \begin{description}
   518 
   519   \item @{command "interpretation"}~@{text "expr \<WHERE> eqns"}
   520   interprets @{text expr} in the theory.  The command generates proof
   521   obligations for the instantiated specifications (assumes and defines
   522   elements).  Once these are discharged by the user, instantiated
   523   facts are added to the theory in a post-processing phase.
   524 
   525   Additional equations, which are unfolded during
   526   post-processing, may be given after the keyword @{keyword "where"}.
   527   This is useful for interpreting concepts introduced through
   528   definitions.  The equations must be proved.
   529 
   530   The command is aware of interpretations already active in the
   531   theory, but does not simplify the goal automatically.  In order to
   532   simplify the proof obligations use methods @{method intro_locales}
   533   or @{method unfold_locales}.  Post-processing is not applied to
   534   facts of interpretations that are already active.  This avoids
   535   duplication of interpreted facts, in particular.  Note that, in the
   536   case of a locale with import, parts of the interpretation may
   537   already be active.  The command will only process facts for new
   538   parts.
   539 
   540   Adding facts to locales has the effect of adding interpreted facts
   541   to the theory for all interpretations as well.  That is,
   542   interpretations dynamically participate in any facts added to
   543   locales.  Note that if a theory inherits additional facts for a
   544   locale through one parent and an interpretation of that locale
   545   through another parent, the additional facts will not be
   546   interpreted.
   547 
   548   \item @{command "interpret"}~@{text "expr \<WHERE> eqns"} interprets
   549   @{text expr} in the proof context and is otherwise similar to
   550   interpretation in theories.  Note that rewrite rules given to
   551   @{command "interpret"} after the @{keyword "where"} keyword should be
   552   explicitly universally quantified.
   553 
   554   \item @{command "sublocale"}~@{text "name \<subseteq> expr \<WHERE>
   555   eqns"}
   556   interprets @{text expr} in the locale @{text name}.  A proof that
   557   the specification of @{text name} implies the specification of
   558   @{text expr} is required.  As in the localized version of the
   559   theorem command, the proof is in the context of @{text name}.  After
   560   the proof obligation has been discharged, the facts of @{text expr}
   561   become part of locale @{text name} as \emph{derived} context
   562   elements and are available when the context @{text name} is
   563   subsequently entered.  Note that, like import, this is dynamic:
   564   facts added to a locale part of @{text expr} after interpretation
   565   become also available in @{text name}.
   566 
   567   Only specification fragments of @{text expr} that are not already
   568   part of @{text name} (be it imported, derived or a derived fragment
   569   of the import) are considered in this process.  This enables
   570   circular interpretations provided that no infinite chains are
   571   generated in the locale hierarchy.
   572 
   573   If interpretations of @{text name} exist in the current theory, the
   574   command adds interpretations for @{text expr} as well, with the same
   575   qualifier, although only for fragments of @{text expr} that are not
   576   interpreted in the theory already.
   577 
   578   Equations given after @{keyword "where"} amend the morphism through
   579   which @{text expr} is interpreted.  This enables to map definitions
   580   from the interpreted locales to entities of @{text name}.  This
   581   feature is experimental.
   582 
   583   \item @{command "print_dependencies"}~@{text "expr"} is useful for
   584   understanding the effect of an interpretation of @{text "expr"}.  It
   585   lists all locale instances for which interpretations would be added
   586   to the current context.  Variant @{command
   587   "print_dependencies"}@{text "!"} prints all locale instances that
   588   would be considered for interpretation, and would be interpreted in
   589   an empty context (that is, without interpretations).
   590 
   591   \item @{command "print_interps"}~@{text "locale"} lists all
   592   interpretations of @{text "locale"} in the current theory or proof
   593   context, including those due to a combination of a @{command
   594   "interpretation"} or @{command "interpret"} and one or several
   595   @{command "sublocale"} declarations.
   596 
   597   \end{description}
   598 
   599   \begin{warn}
   600     Since attributes are applied to interpreted theorems,
   601     interpretation may modify the context of common proof tools, e.g.\
   602     the Simplifier or Classical Reasoner.  As the behavior of such
   603     tools is \emph{not} stable under interpretation morphisms, manual
   604     declarations might have to be added to the target context of the
   605     interpretation to revert such declarations.
   606   \end{warn}
   607 
   608   \begin{warn}
   609     An interpretation in a theory or proof context may subsume previous
   610     interpretations.  This happens if the same specification fragment
   611     is interpreted twice and the instantiation of the second
   612     interpretation is more general than the interpretation of the
   613     first.  The locale package does not attempt to remove subsumed
   614     interpretations.
   615   \end{warn}
   616 *}
   617 
   618 
   619 section {* Classes \label{sec:class} *}
   620 
   621 text {*
   622   A class is a particular locale with \emph{exactly one} type variable
   623   @{text \<alpha>}.  Beyond the underlying locale, a corresponding type class
   624   is established which is interpreted logically as axiomatic type
   625   class \cite{Wenzel:1997:TPHOL} whose logical content are the
   626   assumptions of the locale.  Thus, classes provide the full
   627   generality of locales combined with the commodity of type classes
   628   (notably type-inference).  See \cite{isabelle-classes} for a short
   629   tutorial.
   630 
   631   \begin{matharray}{rcl}
   632     @{command_def "class"} & : & @{text "theory \<rightarrow> local_theory"} \\
   633     @{command_def "instantiation"} & : & @{text "theory \<rightarrow> local_theory"} \\
   634     @{command_def "instance"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   635     @{command "instance"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   636     @{command_def "subclass"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   637     @{command_def "print_classes"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   638     @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   639     @{method_def intro_classes} & : & @{text method} \\
   640   \end{matharray}
   641 
   642   @{rail "
   643     @@{command class} class_spec @'begin'?
   644     ;
   645     class_spec: @{syntax name} '='
   646       ((@{syntax nameref} '+' (@{syntax context_elem}+)) |
   647         @{syntax nameref} | (@{syntax context_elem}+))      
   648     ;
   649     @@{command instantiation} (@{syntax nameref} + @'and') '::' @{syntax arity} @'begin'
   650     ;
   651     @@{command instance} (() | (@{syntax nameref} + @'and') '::' @{syntax arity} |
   652       @{syntax nameref} ('<' | '\<subseteq>') @{syntax nameref} )
   653     ;
   654     @@{command subclass} @{syntax target}? @{syntax nameref}
   655   "}
   656 
   657   \begin{description}
   658 
   659   \item @{command "class"}~@{text "c = superclasses + body"} defines
   660   a new class @{text c}, inheriting from @{text superclasses}.  This
   661   introduces a locale @{text c} with import of all locales @{text
   662   superclasses}.
   663 
   664   Any @{element "fixes"} in @{text body} are lifted to the global
   665   theory level (\emph{class operations} @{text "f\<^sub>1, \<dots>,
   666   f\<^sub>n"} of class @{text c}), mapping the local type parameter
   667   @{text \<alpha>} to a schematic type variable @{text "?\<alpha> :: c"}.
   668 
   669   Likewise, @{element "assumes"} in @{text body} are also lifted,
   670   mapping each local parameter @{text "f :: \<tau>[\<alpha>]"} to its
   671   corresponding global constant @{text "f :: \<tau>[?\<alpha> :: c]"}.  The
   672   corresponding introduction rule is provided as @{text
   673   c_class_axioms.intro}.  This rule should be rarely needed directly
   674   --- the @{method intro_classes} method takes care of the details of
   675   class membership proofs.
   676 
   677   \item @{command "instantiation"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s
   678   \<BEGIN>"} opens a theory target (cf.\ \secref{sec:target}) which
   679   allows to specify class operations @{text "f\<^sub>1, \<dots>, f\<^sub>n"} corresponding
   680   to sort @{text s} at the particular type instance @{text "(\<alpha>\<^sub>1 :: s\<^sub>1,
   681   \<dots>, \<alpha>\<^sub>n :: s\<^sub>n) t"}.  A plain @{command "instance"} command in the
   682   target body poses a goal stating these type arities.  The target is
   683   concluded by an @{command_ref (local) "end"} command.
   684 
   685   Note that a list of simultaneous type constructors may be given;
   686   this corresponds nicely to mutually recursive type definitions, e.g.\
   687   in Isabelle/HOL.
   688 
   689   \item @{command "instance"} in an instantiation target body sets
   690   up a goal stating the type arities claimed at the opening @{command
   691   "instantiation"}.  The proof would usually proceed by @{method
   692   intro_classes}, and then establish the characteristic theorems of
   693   the type classes involved.  After finishing the proof, the
   694   background theory will be augmented by the proven type arities.
   695 
   696   On the theory level, @{command "instance"}~@{text "t :: (s\<^sub>1, \<dots>,
   697   s\<^sub>n)s"} provides a convenient way to instantiate a type class with no
   698   need to specify operations: one can continue with the
   699   instantiation proof immediately.
   700 
   701   \item @{command "subclass"}~@{text c} in a class context for class
   702   @{text d} sets up a goal stating that class @{text c} is logically
   703   contained in class @{text d}.  After finishing the proof, class
   704   @{text d} is proven to be subclass @{text c} and the locale @{text
   705   c} is interpreted into @{text d} simultaneously.
   706 
   707   A weakend form of this is available through a further variant of
   708   @{command instance}:  @{command instance}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} opens
   709   a proof that class @{text "c\<^isub>2"} implies @{text "c\<^isub>1"} without reference
   710   to the underlying locales;  this is useful if the properties to prove
   711   the logical connection are not sufficent on the locale level but on
   712   the theory level.
   713 
   714   \item @{command "print_classes"} prints all classes in the current
   715   theory.
   716 
   717   \item @{command "class_deps"} visualizes all classes and their
   718   subclass relations as a Hasse diagram.
   719 
   720   \item @{method intro_classes} repeatedly expands all class
   721   introduction rules of this theory.  Note that this method usually
   722   needs not be named explicitly, as it is already included in the
   723   default proof step (e.g.\ of @{command "proof"}).  In particular,
   724   instantiation of trivial (syntactic) classes may be performed by a
   725   single ``@{command ".."}'' proof step.
   726 
   727   \end{description}
   728 *}
   729 
   730 
   731 subsection {* The class target *}
   732 
   733 text {*
   734   %FIXME check
   735 
   736   A named context may refer to a locale (cf.\ \secref{sec:target}).
   737   If this locale is also a class @{text c}, apart from the common
   738   locale target behaviour the following happens.
   739 
   740   \begin{itemize}
   741 
   742   \item Local constant declarations @{text "g[\<alpha>]"} referring to the
   743   local type parameter @{text \<alpha>} and local parameters @{text "f[\<alpha>]"}
   744   are accompanied by theory-level constants @{text "g[?\<alpha> :: c]"}
   745   referring to theory-level class operations @{text "f[?\<alpha> :: c]"}.
   746 
   747   \item Local theorem bindings are lifted as are assumptions.
   748 
   749   \item Local syntax refers to local operations @{text "g[\<alpha>]"} and
   750   global operations @{text "g[?\<alpha> :: c]"} uniformly.  Type inference
   751   resolves ambiguities.  In rare cases, manual type annotations are
   752   needed.
   753   
   754   \end{itemize}
   755 *}
   756 
   757 
   758 subsection {* Co-regularity of type classes and arities *}
   759 
   760 text {* The class relation together with the collection of
   761   type-constructor arities must obey the principle of
   762   \emph{co-regularity} as defined below.
   763 
   764   \medskip For the subsequent formulation of co-regularity we assume
   765   that the class relation is closed by transitivity and reflexivity.
   766   Moreover the collection of arities @{text "t :: (\<^vec>s)c"} is
   767   completed such that @{text "t :: (\<^vec>s)c"} and @{text "c \<subseteq> c'"}
   768   implies @{text "t :: (\<^vec>s)c'"} for all such declarations.
   769 
   770   Treating sorts as finite sets of classes (meaning the intersection),
   771   the class relation @{text "c\<^sub>1 \<subseteq> c\<^sub>2"} is extended to sorts as
   772   follows:
   773   \[
   774     @{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"}
   775   \]
   776 
   777   This relation on sorts is further extended to tuples of sorts (of
   778   the same length) in the component-wise way.
   779 
   780   \smallskip Co-regularity of the class relation together with the
   781   arities relation means:
   782   \[
   783     @{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"}
   784   \]
   785   \noindent for all such arities.  In other words, whenever the result
   786   classes of some type-constructor arities are related, then the
   787   argument sorts need to be related in the same way.
   788 
   789   \medskip Co-regularity is a very fundamental property of the
   790   order-sorted algebra of types.  For example, it entails principle
   791   types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}.
   792 *}
   793 
   794 
   795 section {* Unrestricted overloading *}
   796 
   797 text {*
   798   Isabelle/Pure's definitional schemes support certain forms of
   799   overloading (see \secref{sec:consts}).  Overloading means that a
   800   constant being declared as @{text "c :: \<alpha> decl"} may be
   801   defined separately on type instances
   802   @{text "c :: (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n) t decl"}
   803   for each type constructor @{text t}.  At most occassions
   804   overloading will be used in a Haskell-like fashion together with
   805   type classes by means of @{command "instantiation"} (see
   806   \secref{sec:class}).  Sometimes low-level overloading is desirable.
   807   The @{command "overloading"} target provides a convenient view for
   808   end-users.
   809 
   810   \begin{matharray}{rcl}
   811     @{command_def "overloading"} & : & @{text "theory \<rightarrow> local_theory"} \\
   812   \end{matharray}
   813 
   814   @{rail "
   815     @@{command overloading} ( spec + ) @'begin'
   816     ;
   817     spec: @{syntax name} ( '==' | '\<equiv>' ) @{syntax term} ( '(' @'unchecked' ')' )?
   818   "}
   819 
   820   \begin{description}
   821 
   822   \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>"}
   823   opens a theory target (cf.\ \secref{sec:target}) which allows to
   824   specify constants with overloaded definitions.  These are identified
   825   by an explicitly given mapping from variable names @{text "x\<^sub>i"} to
   826   constants @{text "c\<^sub>i"} at particular type instances.  The
   827   definitions themselves are established using common specification
   828   tools, using the names @{text "x\<^sub>i"} as reference to the
   829   corresponding constants.  The target is concluded by @{command
   830   (local) "end"}.
   831 
   832   A @{text "(unchecked)"} option disables global dependency checks for
   833   the corresponding definition, which is occasionally useful for
   834   exotic overloading (see \secref{sec:consts} for a precise description).
   835   It is at the discretion of the user to avoid
   836   malformed theory specifications!
   837 
   838   \end{description}
   839 *}
   840 
   841 
   842 section {* Incorporating ML code \label{sec:ML} *}
   843 
   844 text {*
   845   \begin{matharray}{rcl}
   846     @{command_def "use"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   847     @{command_def "ML"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   848     @{command_def "ML_prf"} & : & @{text "proof \<rightarrow> proof"} \\
   849     @{command_def "ML_val"} & : & @{text "any \<rightarrow>"} \\
   850     @{command_def "ML_command"} & : & @{text "any \<rightarrow>"} \\
   851     @{command_def "setup"} & : & @{text "theory \<rightarrow> theory"} \\
   852     @{command_def "local_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   853     @{command_def "attribute_setup"} & : & @{text "theory \<rightarrow> theory"} \\
   854   \end{matharray}
   855 
   856   @{rail "
   857     @@{command use} @{syntax name}
   858     ;
   859     (@@{command ML} | @@{command ML_prf} | @@{command ML_val} |
   860       @@{command ML_command} | @@{command setup} | @@{command local_setup}) @{syntax text}
   861     ;
   862     @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}?
   863   "}
   864 
   865   \begin{description}
   866 
   867   \item @{command "use"}~@{text "file"} reads and executes ML
   868   commands from @{text "file"}.  The current theory context is passed
   869   down to the ML toplevel and may be modified, using @{ML
   870   "Context.>>"} or derived ML commands.  The file name is checked with
   871   the @{keyword_ref "uses"} dependency declaration given in the theory
   872   header (see also \secref{sec:begin-thy}).
   873 
   874   Top-level ML bindings are stored within the (global or local) theory
   875   context.
   876   
   877   \item @{command "ML"}~@{text "text"} is similar to @{command "use"},
   878   but executes ML commands directly from the given @{text "text"}.
   879   Top-level ML bindings are stored within the (global or local) theory
   880   context.
   881 
   882   \item @{command "ML_prf"} is analogous to @{command "ML"} but works
   883   within a proof context.
   884 
   885   Top-level ML bindings are stored within the proof context in a
   886   purely sequential fashion, disregarding the nested proof structure.
   887   ML bindings introduced by @{command "ML_prf"} are discarded at the
   888   end of the proof.
   889 
   890   \item @{command "ML_val"} and @{command "ML_command"} are diagnostic
   891   versions of @{command "ML"}, which means that the context may not be
   892   updated.  @{command "ML_val"} echos the bindings produced at the ML
   893   toplevel, but @{command "ML_command"} is silent.
   894   
   895   \item @{command "setup"}~@{text "text"} changes the current theory
   896   context by applying @{text "text"}, which refers to an ML expression
   897   of type @{ML_type "theory -> theory"}.  This enables to initialize
   898   any object-logic specific tools and packages written in ML, for
   899   example.
   900 
   901   \item @{command "local_setup"} is similar to @{command "setup"} for
   902   a local theory context, and an ML expression of type @{ML_type
   903   "local_theory -> local_theory"}.  This allows to
   904   invoke local theory specification packages without going through
   905   concrete outer syntax, for example.
   906 
   907   \item @{command "attribute_setup"}~@{text "name = text description"}
   908   defines an attribute in the current theory.  The given @{text
   909   "text"} has to be an ML expression of type
   910   @{ML_type "attribute context_parser"}, cf.\ basic parsers defined in
   911   structure @{ML_struct Args} and @{ML_struct Attrib}.
   912 
   913   In principle, attributes can operate both on a given theorem and the
   914   implicit context, although in practice only one is modified and the
   915   other serves as parameter.  Here are examples for these two cases:
   916 
   917   \end{description}
   918 *}
   919 
   920   attribute_setup my_rule = {*
   921     Attrib.thms >> (fn ths =>
   922       Thm.rule_attribute
   923         (fn context: Context.generic => fn th: thm =>
   924           let val th' = th OF ths
   925           in th' end)) *}
   926 
   927   attribute_setup my_declaration = {*
   928     Attrib.thms >> (fn ths =>
   929       Thm.declaration_attribute
   930         (fn th: thm => fn context: Context.generic =>
   931           let val context' = context
   932           in context' end)) *}
   933 
   934 
   935 section {* Primitive specification elements *}
   936 
   937 subsection {* Type classes and sorts \label{sec:classes} *}
   938 
   939 text {*
   940   \begin{matharray}{rcll}
   941     @{command_def "classes"} & : & @{text "theory \<rightarrow> theory"} \\
   942     @{command_def "classrel"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   943     @{command_def "default_sort"} & : & @{text "local_theory \<rightarrow> local_theory"}
   944   \end{matharray}
   945 
   946   @{rail "
   947     @@{command classes} (@{syntax classdecl} +)
   948     ;
   949     @@{command classrel} (@{syntax nameref} ('<' | '\<subseteq>') @{syntax nameref} + @'and')
   950     ;
   951     @@{command default_sort} @{syntax sort}
   952   "}
   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 "subclass"} and
   964   @{command_ref "instance"} commands (see \secref{sec:class}) provide
   965   a way to introduce proven class relations.
   966 
   967   \item @{command "default_sort"}~@{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   @{class type}, the class of all HOL types.
   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   \end{description}
   981 *}
   982 
   983 
   984 subsection {* Types and type abbreviations \label{sec:types-pure} *}
   985 
   986 text {*
   987   \begin{matharray}{rcll}
   988     @{command_def "type_synonym"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   989     @{command_def "typedecl"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   990     @{command_def "arities"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   991   \end{matharray}
   992 
   993   @{rail "
   994     @@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?)
   995     ;
   996     @@{command typedecl} @{syntax typespec} @{syntax mixfix}?
   997     ;
   998     @@{command arities} (@{syntax nameref} '::' @{syntax arity} +)
   999   "}
  1000 
  1001   \begin{description}
  1002 
  1003   \item @{command "type_synonym"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = \<tau>"}
  1004   introduces a \emph{type synonym} @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} for the
  1005   existing type @{text "\<tau>"}.  Unlike actual type definitions, as are
  1006   available in Isabelle/HOL for example, type synonyms are merely
  1007   syntactic abbreviations without any logical significance.
  1008   Internally, type synonyms are fully expanded.
  1009   
  1010   \item @{command "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} declares a new
  1011   type constructor @{text t}.  If the object-logic defines a base sort
  1012   @{text s}, then the constructor is declared to operate on that, via
  1013   the axiomatic specification @{command arities}~@{text "t :: (s, \<dots>,
  1014   s)s"}.
  1015 
  1016   \item @{command "arities"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n)s"} augments
  1017   Isabelle's order-sorted signature of types by new type constructor
  1018   arities.  This is done axiomatically!  The @{command_ref "instantiation"}
  1019   target (see \secref{sec:class}) provides a way to introduce
  1020   proven type arities.
  1021 
  1022   \end{description}
  1023 *}
  1024 
  1025 
  1026 subsection {* Constants and definitions \label{sec:consts} *}
  1027 
  1028 text {*
  1029   Definitions essentially express abbreviations within the logic.  The
  1030   simplest form of a definition is @{text "c :: \<sigma> \<equiv> t"}, where @{text
  1031   c} is a newly declared constant.  Isabelle also allows derived forms
  1032   where the arguments of @{text c} appear on the left, abbreviating a
  1033   prefix of @{text \<lambda>}-abstractions, e.g.\ @{text "c \<equiv> \<lambda>x y. t"} may be
  1034   written more conveniently as @{text "c x y \<equiv> t"}.  Moreover,
  1035   definitions may be weakened by adding arbitrary pre-conditions:
  1036   @{text "A \<Longrightarrow> c x y \<equiv> t"}.
  1037 
  1038   \medskip The built-in well-formedness conditions for definitional
  1039   specifications are:
  1040 
  1041   \begin{itemize}
  1042 
  1043   \item Arguments (on the left-hand side) must be distinct variables.
  1044 
  1045   \item All variables on the right-hand side must also appear on the
  1046   left-hand side.
  1047 
  1048   \item All type variables on the right-hand side must also appear on
  1049   the left-hand side; this prohibits @{text "0 :: nat \<equiv> length ([] ::
  1050   \<alpha> list)"} for example.
  1051 
  1052   \item The definition must not be recursive.  Most object-logics
  1053   provide definitional principles that can be used to express
  1054   recursion safely.
  1055 
  1056   \end{itemize}
  1057 
  1058   The right-hand side of overloaded definitions may mention overloaded constants
  1059   recursively at type instances corresponding to the immediate
  1060   argument types @{text "\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n"}.  Incomplete
  1061   specification patterns impose global constraints on all occurrences,
  1062   e.g.\ @{text "d :: \<alpha> \<times> \<alpha>"} on the left-hand side means that all
  1063   corresponding occurrences on some right-hand side need to be an
  1064   instance of this, general @{text "d :: \<alpha> \<times> \<beta>"} will be disallowed.
  1065 
  1066   \begin{matharray}{rcl}
  1067     @{command_def "consts"} & : & @{text "theory \<rightarrow> theory"} \\
  1068     @{command_def "defs"} & : & @{text "theory \<rightarrow> theory"} \\
  1069   \end{matharray}
  1070 
  1071   @{rail "
  1072     @@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +)
  1073     ;
  1074     @@{command defs} opt? (@{syntax axmdecl} @{syntax prop} +)
  1075     ;
  1076     opt: '(' @'unchecked'? @'overloaded'? ')'
  1077   "}
  1078 
  1079   \begin{description}
  1080 
  1081   \item @{command "consts"}~@{text "c :: \<sigma>"} declares constant @{text
  1082   c} to have any instance of type scheme @{text \<sigma>}.  The optional
  1083   mixfix annotations may attach concrete syntax to the constants
  1084   declared.
  1085   
  1086   \item @{command "defs"}~@{text "name: eqn"} introduces @{text eqn}
  1087   as a definitional axiom for some existing constant.
  1088   
  1089   The @{text "(unchecked)"} option disables global dependency checks
  1090   for this definition, which is occasionally useful for exotic
  1091   overloading.  It is at the discretion of the user to avoid malformed
  1092   theory specifications!
  1093   
  1094   The @{text "(overloaded)"} option declares definitions to be
  1095   potentially overloaded.  Unless this option is given, a warning
  1096   message would be issued for any definitional equation with a more
  1097   special type than that of the corresponding constant declaration.
  1098   
  1099   \end{description}
  1100 *}
  1101 
  1102 
  1103 section {* Axioms and theorems \label{sec:axms-thms} *}
  1104 
  1105 text {*
  1106   \begin{matharray}{rcll}
  1107     @{command_def "axioms"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
  1108     @{command_def "lemmas"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1109     @{command_def "theorems"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1110   \end{matharray}
  1111 
  1112   @{rail "
  1113     @@{command axioms} (@{syntax axmdecl} @{syntax prop} +)
  1114     ;
  1115     (@@{command lemmas} | @@{command theorems}) @{syntax target}? \\
  1116       (@{syntax thmdef}? @{syntax thmrefs} + @'and')
  1117       (@'for' (@{syntax vars} + @'and'))?
  1118   "}
  1119 
  1120   \begin{description}
  1121   
  1122   \item @{command "axioms"}~@{text "a: \<phi>"} introduces arbitrary
  1123   statements as axioms of the meta-logic.  In fact, axioms are
  1124   ``axiomatic theorems'', and may be referred later just as any other
  1125   theorem.
  1126   
  1127   Axioms are usually only introduced when declaring new logical
  1128   systems.  Everyday work is typically done the hard way, with proper
  1129   definitions and proven theorems.
  1130   
  1131   \item @{command "lemmas"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"}~@{keyword_def
  1132   "for"}~@{text "x\<^sub>1 \<dots> x\<^sub>m"} evaluates given facts (with attributes) in
  1133   the current context, which may be augmented by local variables.
  1134   Results are standardized before being stored, i.e.\ schematic
  1135   variables are renamed to enforce index @{text "0"} uniformly.
  1136 
  1137   \item @{command "theorems"} is the same as @{command "lemmas"}, but
  1138   marks the result as a different kind of facts.
  1139 
  1140   \end{description}
  1141 *}
  1142 
  1143 
  1144 section {* Oracles *}
  1145 
  1146 text {* Oracles allow Isabelle to take advantage of external reasoners
  1147   such as arithmetic decision procedures, model checkers, fast
  1148   tautology checkers or computer algebra systems.  Invoked as an
  1149   oracle, an external reasoner can create arbitrary Isabelle theorems.
  1150 
  1151   It is the responsibility of the user to ensure that the external
  1152   reasoner is as trustworthy as the application requires.  Another
  1153   typical source of errors is the linkup between Isabelle and the
  1154   external tool, not just its concrete implementation, but also the
  1155   required translation between two different logical environments.
  1156 
  1157   Isabelle merely guarantees well-formedness of the propositions being
  1158   asserted, and records within the internal derivation object how
  1159   presumed theorems depend on unproven suppositions.
  1160 
  1161   \begin{matharray}{rcll}
  1162     @{command_def "oracle"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
  1163   \end{matharray}
  1164 
  1165   @{rail "
  1166     @@{command oracle} @{syntax name} '=' @{syntax text}
  1167   "}
  1168 
  1169   \begin{description}
  1170 
  1171   \item @{command "oracle"}~@{text "name = text"} turns the given ML
  1172   expression @{text "text"} of type @{ML_text "'a -> cterm"} into an
  1173   ML function of type @{ML_text "'a -> thm"}, which is bound to the
  1174   global identifier @{ML_text name}.  This acts like an infinitary
  1175   specification of axioms!  Invoking the oracle only works within the
  1176   scope of the resulting theory.
  1177 
  1178   \end{description}
  1179 
  1180   See @{file "~~/src/HOL/ex/Iff_Oracle.thy"} for a worked example of
  1181   defining a new primitive rule as oracle, and turning it into a proof
  1182   method.
  1183 *}
  1184 
  1185 
  1186 section {* Name spaces *}
  1187 
  1188 text {*
  1189   \begin{matharray}{rcl}
  1190     @{command_def "hide_class"} & : & @{text "theory \<rightarrow> theory"} \\
  1191     @{command_def "hide_type"} & : & @{text "theory \<rightarrow> theory"} \\
  1192     @{command_def "hide_const"} & : & @{text "theory \<rightarrow> theory"} \\
  1193     @{command_def "hide_fact"} & : & @{text "theory \<rightarrow> theory"} \\
  1194   \end{matharray}
  1195 
  1196   @{rail "
  1197     ( @{command hide_class} | @{command hide_type} |
  1198       @{command hide_const} | @{command hide_fact} ) ('(' @'open' ')')? (@{syntax nameref} + )
  1199   "}
  1200 
  1201   Isabelle organizes any kind of name declarations (of types,
  1202   constants, theorems etc.) by separate hierarchically structured name
  1203   spaces.  Normally the user does not have to control the behavior of
  1204   name spaces by hand, yet the following commands provide some way to
  1205   do so.
  1206 
  1207   \begin{description}
  1208 
  1209   \item @{command "hide_class"}~@{text names} fully removes class
  1210   declarations from a given name space; with the @{text "(open)"}
  1211   option, only the base name is hidden.
  1212 
  1213   Note that hiding name space accesses has no impact on logical
  1214   declarations --- they remain valid internally.  Entities that are no
  1215   longer accessible to the user are printed with the special qualifier
  1216   ``@{text "??"}'' prefixed to the full internal name.
  1217 
  1218   \item @{command "hide_type"}, @{command "hide_const"}, and @{command
  1219   "hide_fact"} are similar to @{command "hide_class"}, but hide types,
  1220   constants, and facts, respectively.
  1221   
  1222   \end{description}
  1223 *}
  1224 
  1225 end