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