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