src/Doc/Isar_Ref/Spec.thy
author wenzelm
Sat, 27 Feb 2021 16:33:16 +0100
changeset 73313 8ae2f8ebc373
parent 72804 d7393c35aa5d
child 73757 cb933ba9ecfe
permissions -rw-r--r--
discontinued somewhat pointless "integrity test of build_history": it fails right now, but also failed to expose spurious incompatibilities when it was working;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61656
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61606
diff changeset
     1
(*:maxLineLen=78:*)
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61606
diff changeset
     2
26869
3bc332135aa7 added chapters for "Specifications" and "Proofs";
wenzelm
parents:
diff changeset
     3
theory Spec
63531
847eefdca90d clarified imports;
wenzelm
parents: 63285
diff changeset
     4
  imports Main Base
26869
3bc332135aa7 added chapters for "Specifications" and "Proofs";
wenzelm
parents:
diff changeset
     5
begin
3bc332135aa7 added chapters for "Specifications" and "Proofs";
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     7
chapter \<open>Specifications\<close>
26869
3bc332135aa7 added chapters for "Specifications" and "Proofs";
wenzelm
parents:
diff changeset
     8
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
     9
text \<open>
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    10
  The Isabelle/Isar theory format integrates specifications and proofs, with
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    11
  support for interactive development by continuous document editing. There is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    12
  a separate document preparation system (see \chref{ch:document-prep}), for
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    13
  typesetting formal developments together with informal text. The resulting
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    14
  hyper-linked PDF documents can be used both for WWW presentation and printed
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    15
  copies.
29745
fe221f1d8976 added introduction;
wenzelm
parents: 29706
diff changeset
    16
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    17
  The Isar proof language (see \chref{ch:proofs}) is embedded into the theory
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    18
  language as a proper sub-language. Proof mode is entered by stating some
61665
wenzelm
parents: 61664
diff changeset
    19
  \<^theory_text>\<open>theorem\<close> or \<^theory_text>\<open>lemma\<close> at the theory level, and left again with the final
wenzelm
parents: 61664
diff changeset
    20
  conclusion (e.g.\ via \<^theory_text>\<open>qed\<close>).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    21
\<close>
29745
fe221f1d8976 added introduction;
wenzelm
parents: 29706
diff changeset
    22
fe221f1d8976 added introduction;
wenzelm
parents: 29706
diff changeset
    23
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    24
section \<open>Defining theories \label{sec:begin-thy}\<close>
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    25
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    26
text \<open>
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    27
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    28
    @{command_def "theory"} & : & \<open>toplevel \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    29
    @{command_def (global) "end"} & : & \<open>theory \<rightarrow> toplevel\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    30
    @{command_def "thy_deps"}\<open>\<^sup>*\<close> & : & \<open>theory \<rightarrow>\<close> \\
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    31
  \end{matharray}
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    32
59781
wenzelm
parents: 59422
diff changeset
    33
  Isabelle/Isar theories are defined via theory files, which consist of an
wenzelm
parents: 59422
diff changeset
    34
  outermost sequence of definition--statement--proof elements. Some
61665
wenzelm
parents: 61664
diff changeset
    35
  definitions are self-sufficient (e.g.\ \<^theory_text>\<open>fun\<close> in Isabelle/HOL), with
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    36
  foundational proofs performed internally. Other definitions require an
61665
wenzelm
parents: 61664
diff changeset
    37
  explicit proof as justification (e.g.\ \<^theory_text>\<open>function\<close> and \<^theory_text>\<open>termination\<close> in
wenzelm
parents: 61664
diff changeset
    38
  Isabelle/HOL). Plain statements like \<^theory_text>\<open>theorem\<close> or \<^theory_text>\<open>lemma\<close> are merely a
wenzelm
parents: 61664
diff changeset
    39
  special case of that, defining a theorem from a given proposition and its
wenzelm
parents: 61664
diff changeset
    40
  proof.
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    41
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    42
  The theory body may be sub-structured by means of \<^emph>\<open>local theory targets\<close>,
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
    43
  such as \<^theory_text>\<open>locale\<close> and \<^theory_text>\<open>class\<close>. It is also possible to use \<^theory_text>\<open>context begin \<dots>
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
    44
  end\<close> blocks to delimited a local theory context: a \<^emph>\<open>named context\<close> to
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
    45
  augment a locale or class specification, or an \<^emph>\<open>unnamed context\<close> to refer
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
    46
  to local parameters and assumptions that are discharged later. See
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
    47
  \secref{sec:target} for more details.
28745
146d570e12b5 reworked "Defining Theories";
wenzelm
parents: 28728
diff changeset
    48
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
    49
  \<^medskip>
61665
wenzelm
parents: 61664
diff changeset
    50
  A theory is commenced by the \<^theory_text>\<open>theory\<close> command, which indicates imports of
wenzelm
parents: 61664
diff changeset
    51
  previous theories, according to an acyclic foundational order. Before the
wenzelm
parents: 61664
diff changeset
    52
  initial \<^theory_text>\<open>theory\<close> command, there may be optional document header material
wenzelm
parents: 61664
diff changeset
    53
  (like \<^theory_text>\<open>section\<close> or \<^theory_text>\<open>text\<close>, see \secref{sec:markup}). The document header
wenzelm
parents: 61664
diff changeset
    54
  is outside of the formal theory context, though.
59781
wenzelm
parents: 59422
diff changeset
    55
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    56
  A theory is concluded by a final @{command (global) "end"} command, one that
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    57
  does not belong to a local theory target. No further commands may follow
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    58
  such a global @{command (global) "end"}.
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    59
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
    60
  \<^rail>\<open>
66916
aca50a1572c5 more documentation;
wenzelm
parents: 66757
diff changeset
    61
    @@{command theory} @{syntax system_name}
aca50a1572c5 more documentation;
wenzelm
parents: 66757
diff changeset
    62
      @'imports' (@{syntax system_name} +) \<newline>
63580
7f06347a5013 clarified: 'imports' is de-facto mandatory;
wenzelm
parents: 63579
diff changeset
    63
      keywords? abbrevs? @'begin'
47114
7c9e31ffcd9e updated theory header syntax and related details;
wenzelm
parents: 46999
diff changeset
    64
    ;
50128
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49243
diff changeset
    65
    keywords: @'keywords' (keyword_decls + @'and')
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49243
diff changeset
    66
    ;
63579
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
    67
    keyword_decls: (@{syntax string} +) ('::' @{syntax name} @{syntax tags})?
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
    68
    ;
67722
012f1e8a1209 allow multiple entries of and_list (on both sides);
wenzelm
parents: 67013
diff changeset
    69
    abbrevs: @'abbrevs' (((text+) '=' (text+)) + @'and')
60093
c48d536231fe clarified thy_deps;
wenzelm
parents: 60091
diff changeset
    70
    ;
c48d536231fe clarified thy_deps;
wenzelm
parents: 60091
diff changeset
    71
    @@{command thy_deps} (thy_bounds thy_bounds?)?
c48d536231fe clarified thy_deps;
wenzelm
parents: 60091
diff changeset
    72
    ;
c48d536231fe clarified thy_deps;
wenzelm
parents: 60091
diff changeset
    73
    thy_bounds: @{syntax name} | '(' (@{syntax name} + @'|') ')'
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
    74
  \<close>
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
    75
61665
wenzelm
parents: 61664
diff changeset
    76
  \<^descr> \<^theory_text>\<open>theory A imports B\<^sub>1 \<dots> B\<^sub>n begin\<close> starts a new theory \<open>A\<close> based on the
wenzelm
parents: 61664
diff changeset
    77
  merge of existing theories \<open>B\<^sub>1 \<dots> B\<^sub>n\<close>. Due to the possibility to import
wenzelm
parents: 61664
diff changeset
    78
  more than one ancestor, the resulting theory structure of an Isabelle
wenzelm
parents: 61664
diff changeset
    79
  session forms a directed acyclic graph (DAG). Isabelle takes care that
wenzelm
parents: 61664
diff changeset
    80
  sources contributing to the development graph are always up-to-date: changed
wenzelm
parents: 61664
diff changeset
    81
  files are automatically rechecked whenever a theory header specification is
wenzelm
parents: 61664
diff changeset
    82
  processed.
47114
7c9e31ffcd9e updated theory header syntax and related details;
wenzelm
parents: 46999
diff changeset
    83
59781
wenzelm
parents: 59422
diff changeset
    84
  Empty imports are only allowed in the bootstrap process of the special
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
    85
  theory \<^theory>\<open>Pure\<close>, which is the start of any other formal development
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    86
  based on Isabelle. Regular user theories usually refer to some more complex
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
    87
  entry point, such as theory \<^theory>\<open>Main\<close> in Isabelle/HOL.
59781
wenzelm
parents: 59422
diff changeset
    88
63579
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
    89
  The @{keyword_def "keywords"} specification declares outer syntax
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    90
  (\chref{ch:outer-syntax}) that is introduced in this theory later on (rare
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    91
  in end-user applications). Both minor keywords and major keywords of the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    92
  Isar command language need to be specified, in order to make parsing of
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    93
  proof documents work properly. Command keywords need to be classified
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    94
  according to their structural role in the formal text. Examples may be seen
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    95
  in Isabelle/HOL sources itself, such as @{keyword "keywords"}~\<^verbatim>\<open>"typedef"\<close>
69913
ca515cf61651 more specific keyword kinds;
wenzelm
parents: 69597
diff changeset
    96
  \<open>:: thy_goal_defn\<close> or @{keyword "keywords"}~\<^verbatim>\<open>"datatype"\<close> \<open>:: thy_defn\<close> for
ca515cf61651 more specific keyword kinds;
wenzelm
parents: 69597
diff changeset
    97
  theory-level definitions with and without proof, respectively. Additional
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
    98
  @{syntax tags} provide defaults for document preparation
69962
82e945d472d5 documentation of document markers and re-interpreted command tags;
wenzelm
parents: 69913
diff changeset
    99
  (\secref{sec:document-markers}).
50128
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49243
diff changeset
   100
63579
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
   101
  The @{keyword_def "abbrevs"} specification declares additional abbreviations
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
   102
  for syntactic completion. The default for a new keyword is just its name,
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
   103
  but completion may be avoided by defining @{keyword "abbrevs"} with empty
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 63531
diff changeset
   104
  text.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   105
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   106
  \<^descr> @{command (global) "end"} concludes the current theory definition. Note
61665
wenzelm
parents: 61664
diff changeset
   107
  that some other commands, e.g.\ local theory targets \<^theory_text>\<open>locale\<close> or \<^theory_text>\<open>class\<close>
wenzelm
parents: 61664
diff changeset
   108
  may involve a \<^theory_text>\<open>begin\<close> that needs to be matched by @{command (local) "end"},
wenzelm
parents: 61664
diff changeset
   109
  according to the usual rules for nested blocks.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   110
61665
wenzelm
parents: 61664
diff changeset
   111
  \<^descr> \<^theory_text>\<open>thy_deps\<close> visualizes the theory hierarchy as a directed acyclic graph.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   112
  By default, all imported theories are shown. This may be restricted by
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   113
  specifying bounds wrt. the theory inclusion relation.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   114
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   115
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   116
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   117
section \<open>Local theory targets \label{sec:target}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   118
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   119
text \<open>
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   120
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   121
    @{command_def "context"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   122
    @{command_def (local) "end"} & : & \<open>local_theory \<rightarrow> theory\<close> \\
59926
003dbac78546 some explanation of 'private';
wenzelm
parents: 59917
diff changeset
   123
    @{keyword_def "private"} \\
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   124
    @{keyword_def "qualified"} \\
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   125
  \end{matharray}
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   126
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   127
  A local theory target is a specification context that is managed separately
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   128
  within the enclosing theory. Contexts may introduce parameters (fixed
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   129
  variables) and assumptions (hypotheses). Definitions and theorems depending
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   130
  on the context may be added incrementally later on.
47482
a83b25e5bad3 some coverage of unnamed contexts, which can be nested within other targets;
wenzelm
parents: 47114
diff changeset
   131
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   132
  \<^emph>\<open>Named contexts\<close> refer to locales (cf.\ \secref{sec:locale}) or type
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   133
  classes (cf.\ \secref{sec:class}); the name ``\<open>-\<close>'' signifies the global
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   134
  theory context.
47482
a83b25e5bad3 some coverage of unnamed contexts, which can be nested within other targets;
wenzelm
parents: 47114
diff changeset
   135
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   136
  \<^emph>\<open>Unnamed contexts\<close> may introduce additional parameters and assumptions, and
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   137
  results produced in the context are generalized accordingly. Such auxiliary
61665
wenzelm
parents: 61664
diff changeset
   138
  contexts may be nested within other targets, like \<^theory_text>\<open>locale\<close>, \<^theory_text>\<open>class\<close>,
wenzelm
parents: 61664
diff changeset
   139
  \<^theory_text>\<open>instantiation\<close>, \<^theory_text>\<open>overloading\<close>.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   140
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   141
  \<^rail>\<open>
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   142
    @@{command context} @{syntax name} @{syntax_ref "opening"}? @'begin'
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   143
    ;
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   144
    @@{command context} @{syntax_ref "includes"}? (@{syntax context_elem} * ) @'begin'
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   145
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   146
    @{syntax_def target}: '(' @'in' @{syntax name} ')'
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   147
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   148
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   149
  \<^descr> \<^theory_text>\<open>context c bundles begin\<close> opens a named context, by recommencing an existing
61665
wenzelm
parents: 61664
diff changeset
   150
  locale or class \<open>c\<close>. Note that locale and class definitions allow to include
wenzelm
parents: 61664
diff changeset
   151
  the \<^theory_text>\<open>begin\<close> keyword as well, in order to continue the local theory
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   152
  immediately after the initial specification.  Optionally given
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   153
  \<open>bundles\<close> only take effect in the surface context within the \<^theory_text>\<open>begin\<close> /
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   154
  \<^theory_text>\<open>end\<close> block.
47482
a83b25e5bad3 some coverage of unnamed contexts, which can be nested within other targets;
wenzelm
parents: 47114
diff changeset
   155
61665
wenzelm
parents: 61664
diff changeset
   156
  \<^descr> \<^theory_text>\<open>context bundles elements begin\<close> opens an unnamed context, by extending
wenzelm
parents: 61664
diff changeset
   157
  the enclosing global or local theory target by the given declaration bundles
wenzelm
parents: 61664
diff changeset
   158
  (\secref{sec:bundle}) and context elements (\<^theory_text>\<open>fixes\<close>, \<^theory_text>\<open>assumes\<close> etc.). This
wenzelm
parents: 61664
diff changeset
   159
  means any results stemming from definitions and proofs in the extended
wenzelm
parents: 61664
diff changeset
   160
  context will be exported into the enclosing target by lifting over extra
wenzelm
parents: 61664
diff changeset
   161
  parameters and premises.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   162
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   163
  \<^descr> @{command (local) "end"} concludes the current local theory, according to
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   164
  the nesting of contexts. Note that a global @{command (global) "end"} has a
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   165
  different meaning: it concludes the theory itself (\secref{sec:begin-thy}).
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   166
61665
wenzelm
parents: 61664
diff changeset
   167
  \<^descr> \<^theory_text>\<open>private\<close> or \<^theory_text>\<open>qualified\<close> may be given as modifiers before any local
wenzelm
parents: 61664
diff changeset
   168
  theory command. This restricts name space accesses to the local scope, as
wenzelm
parents: 61664
diff changeset
   169
  determined by the enclosing \<^theory_text>\<open>context begin \<dots> end\<close> block. Outside its scope,
wenzelm
parents: 61664
diff changeset
   170
  a \<^theory_text>\<open>private\<close> name is inaccessible, and a \<^theory_text>\<open>qualified\<close> name is only
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   171
  accessible with some qualification.
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59926
diff changeset
   172
61665
wenzelm
parents: 61664
diff changeset
   173
  Neither a global \<^theory_text>\<open>theory\<close> nor a \<^theory_text>\<open>locale\<close> target provides a local scope by
wenzelm
parents: 61664
diff changeset
   174
  itself: an extra unnamed context is required to use \<^theory_text>\<open>private\<close> or
wenzelm
parents: 61664
diff changeset
   175
  \<^theory_text>\<open>qualified\<close> here.
59926
003dbac78546 some explanation of 'private';
wenzelm
parents: 59917
diff changeset
   176
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   177
  \<^descr> \<open>(\<close>@{keyword_def "in"}~\<open>c)\<close> given after any local theory command specifies
61665
wenzelm
parents: 61664
diff changeset
   178
  an immediate target, e.g.\ ``\<^theory_text>\<open>definition (in c)\<close>'' or
wenzelm
parents: 61664
diff changeset
   179
  ``\<^theory_text>\<open>theorem (in c)\<close>''. This works both in a local or global theory context;
wenzelm
parents: 61664
diff changeset
   180
  the current target context will be suspended for this command only. Note
wenzelm
parents: 61664
diff changeset
   181
  that ``\<^theory_text>\<open>(in -)\<close>'' will always produce a global result independently of the
wenzelm
parents: 61664
diff changeset
   182
  current target context.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   183
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   184
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   185
  Any specification element that operates on \<open>local_theory\<close> according to this
61665
wenzelm
parents: 61664
diff changeset
   186
  manual implicitly allows the above target syntax \<^theory_text>\<open>(in c)\<close>, but individual
wenzelm
parents: 61664
diff changeset
   187
  syntax diagrams omit that aspect for clarity.
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 59781
diff changeset
   188
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   189
  \<^medskip>
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   190
  The exact meaning of results produced within a local theory context depends
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   191
  on the underlying target infrastructure (locale, type class etc.). The
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   192
  general idea is as follows, considering a context named \<open>c\<close> with parameter
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   193
  \<open>x\<close> and assumption \<open>A[x]\<close>.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   194
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   195
  Definitions are exported by introducing a global version with additional
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   196
  arguments; a syntactic abbreviation links the long form with the abstract
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   197
  version of the target context. For example, \<open>a \<equiv> t[x]\<close> becomes \<open>c.a ?x \<equiv>
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   198
  t[?x]\<close> at the theory level (for arbitrary \<open>?x\<close>), together with a local
71567
9a29e883a934 tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents: 71166
diff changeset
   199
  abbreviation \<open>a \<equiv> c.a x\<close> in the target context (for the fixed parameter
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   200
  \<open>x\<close>).
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   201
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   202
  Theorems are exported by discharging the assumptions and generalizing the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   203
  parameters of the context. For example, \<open>a: B[x]\<close> becomes \<open>c.a: A[?x] \<Longrightarrow>
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   204
  B[?x]\<close>, again for arbitrary \<open>?x\<close>.
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 59781
diff changeset
   205
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   206
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   207
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   208
section \<open>Bundled declarations \label{sec:bundle}\<close>
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   209
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   210
text \<open>
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   211
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   212
    @{command_def "bundle"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
63273
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   213
    @{command "bundle"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
61494
63b18f758874 proper spaces around @{text};
wenzelm
parents: 61493
diff changeset
   214
    @{command_def "print_bundles"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   215
    @{command_def "include"} & : & \<open>proof(state) \<rightarrow> proof(state)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   216
    @{command_def "including"} & : & \<open>proof(prove) \<rightarrow> proof(prove)\<close> \\
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   217
    @{keyword_def "includes"} & : & syntax \\
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   218
  \end{matharray}
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   219
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   220
  The outer syntax of fact expressions (\secref{sec:syn-att}) involves
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   221
  theorems and attributes, which are evaluated in the context and applied to
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   222
  it. Attributes may declare theorems to the context, as in \<open>this_rule [intro]
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   223
  that_rule [elim]\<close> for example. Configuration options (\secref{sec:config})
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   224
  are special declaration attributes that operate on the context without a
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   225
  theorem, as in \<open>[[show_types = false]]\<close> for example.
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   226
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   227
  Expressions of this form may be defined as \<^emph>\<open>bundled declarations\<close> in the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   228
  context, and included in other situations later on. Including declaration
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   229
  bundles augments a local context casually without logical dependencies,
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   230
  which is in contrast to locales and locale interpretation
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   231
  (\secref{sec:locale}).
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   232
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   233
  \<^rail>\<open>
63273
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   234
    @@{command bundle} @{syntax name}
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   235
      ( '=' @{syntax thms} @{syntax for_fixes} | @'begin')
63272
6d8a67a77bad documentation;
wenzelm
parents: 63182
diff changeset
   236
    ;
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   237
    @@{command print_bundles} ('!'?)
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   238
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   239
    (@@{command include} | @@{command including}) (@{syntax name}+)
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   240
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   241
    @{syntax_def "includes"}: @'includes' (@{syntax name}+)
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   242
    ;
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   243
    @{syntax_def "opening"}: @'opening' (@{syntax name}+)
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   244
    ;
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   245
    @@{command unbundle} (@{syntax name}+)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   246
  \<close>
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   247
61665
wenzelm
parents: 61664
diff changeset
   248
  \<^descr> \<^theory_text>\<open>bundle b = decls\<close> defines a bundle of declarations in the current
wenzelm
parents: 61664
diff changeset
   249
  context. The RHS is similar to the one of the \<^theory_text>\<open>declare\<close> command. Bundles
wenzelm
parents: 61664
diff changeset
   250
  defined in local theory targets are subject to transformations via
wenzelm
parents: 61664
diff changeset
   251
  morphisms, when moved into different application contexts; this works
wenzelm
parents: 61664
diff changeset
   252
  analogously to any other local theory specification.
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   253
63273
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   254
  \<^descr> \<^theory_text>\<open>bundle b begin body end\<close> defines a bundle of declarations from the
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   255
  \<open>body\<close> of local theory specifications. It may consist of commands that are
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   256
  technically equivalent to \<^theory_text>\<open>declare\<close> or \<^theory_text>\<open>declaration\<close>, which also includes
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   257
  \<^theory_text>\<open>notation\<close>, for example. Named fact declarations like ``\<^theory_text>\<open>lemmas a [simp] =
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   258
  b\<close>'' or ``\<^theory_text>\<open>lemma a [simp]: B \<proof>\<close>'' are also admitted, but the name
302daf918966 prefer hybrid 'bundle' command;
wenzelm
parents: 63272
diff changeset
   259
  bindings are not recorded in the bundle.
63272
6d8a67a77bad documentation;
wenzelm
parents: 63182
diff changeset
   260
61665
wenzelm
parents: 61664
diff changeset
   261
  \<^descr> \<^theory_text>\<open>print_bundles\<close> prints the named bundles that are available in the
wenzelm
parents: 61664
diff changeset
   262
  current context; the ``\<open>!\<close>'' option indicates extra verbosity.
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   263
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   264
  \<^descr> \<^theory_text>\<open>include b\<^sub>1 \<dots> b\<^sub>n\<close> activates the declarations from the given bundles
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   265
  in a proof body (forward mode). This is analogous to \<^theory_text>\<open>note\<close>
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   266
  (\secref{sec:proof-facts}) with the expanded bundles.
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   267
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   268
  \<^descr> \<^theory_text>\<open>including b\<^sub>1 \<dots> b\<^sub>n\<close> is similar to \<^theory_text>\<open>include\<close>, but works in proof refinement
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   269
  (backward mode). This is analogous to \<^theory_text>\<open>using\<close> (\secref{sec:proof-facts})
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   270
  with the expanded bundles.
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   271
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   272
  \<^descr> \<^theory_text>\<open>includes b\<^sub>1 \<dots> b\<^sub>n\<close> is similar to \<^theory_text>\<open>include\<close>, but applies to a
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   273
  confined specification context: unnamed \<^theory_text>\<open>context\<close>s and
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   274
  long statements of \<^theory_text>\<open>theorem\<close>.
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   275
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   276
  \<^descr> \<^theory_text>\<open>opening b\<^sub>1 \<dots> b\<^sub>n\<close> is similar to \<^theory_text>\<open>includes\<close>, but applies to
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   277
  a named specification context: \<^theory_text>\<open>locale\<close>s, \<^theory_text>\<open>class\<close>es and
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   278
  named \<^theory_text>\<open>context\<close>s. The effect is confined to the surface context within the
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   279
  specification block itself and the corresponding \<^theory_text>\<open>begin\<close> / \<^theory_text>\<open>end\<close> block.
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   280
63282
7c509ddf29a5 added command 'unbundle';
wenzelm
parents: 63273
diff changeset
   281
  \<^descr> \<^theory_text>\<open>unbundle b\<^sub>1 \<dots> b\<^sub>n\<close> activates the declarations from the given bundles in
7c509ddf29a5 added command 'unbundle';
wenzelm
parents: 63273
diff changeset
   282
  the current local theory context. This is analogous to \<^theory_text>\<open>lemmas\<close>
7c509ddf29a5 added command 'unbundle';
wenzelm
parents: 63273
diff changeset
   283
  (\secref{sec:theorems}) with the expanded bundles.
7c509ddf29a5 added command 'unbundle';
wenzelm
parents: 63273
diff changeset
   284
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   285
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   286
  Here is an artificial example of bundling various configuration options:
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   287
\<close>
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   288
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59901
diff changeset
   289
(*<*)experiment begin(*>*)
53536
69c943125fd3 do not expose internal flags to attribute name space;
wenzelm
parents: 53370
diff changeset
   290
bundle trace = [[simp_trace, linarith_trace, metis_trace, smt_trace]]
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   291
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   292
lemma "x = x"
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   293
  including trace by metis
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59901
diff changeset
   294
(*<*)end(*>*)
47484
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   295
e94cc23d434a some coverage of bundled declarations;
wenzelm
parents: 47483
diff changeset
   296
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   297
section \<open>Term definitions \label{sec:term-definitions}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   298
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   299
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   300
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   301
    @{command_def "definition"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   302
    @{attribute_def "defn"} & : & \<open>attribute\<close> \\
61494
63b18f758874 proper spaces around @{text};
wenzelm
parents: 61493
diff changeset
   303
    @{command_def "print_defn_rules"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   304
    @{command_def "abbreviation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
61494
63b18f758874 proper spaces around @{text};
wenzelm
parents: 61493
diff changeset
   305
    @{command_def "print_abbrevs"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   306
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   307
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   308
  Term definitions may either happen within the logic (as equational axioms of
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   309
  a certain form (see also \secref{sec:overloading}), or outside of it as
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   310
  rewrite system on abstract syntax. The second form is called
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   311
  ``abbreviation''.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   312
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   313
  \<^rail>\<open>
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   314
    @@{command definition} decl? definition
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   315
    ;
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   316
    @@{command abbreviation} @{syntax mode}? decl? abbreviation
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   317
    ;
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   318
    @@{command print_abbrevs} ('!'?)
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   319
    ;
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   320
    decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? @'where'
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   321
    ;
63182
b065b4833092 allow 'for' fixes for multi_specs;
wenzelm
parents: 63180
diff changeset
   322
    definition: @{syntax thmdecl}? @{syntax prop}
b065b4833092 allow 'for' fixes for multi_specs;
wenzelm
parents: 63180
diff changeset
   323
      @{syntax spec_prems} @{syntax for_fixes}
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   324
    ;
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63178
diff changeset
   325
    abbreviation: @{syntax prop} @{syntax for_fixes}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   326
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   327
61665
wenzelm
parents: 61664
diff changeset
   328
  \<^descr> \<^theory_text>\<open>definition c where eq\<close> produces an internal definition \<open>c \<equiv> t\<close> according
wenzelm
parents: 61664
diff changeset
   329
  to the specification given as \<open>eq\<close>, which is then turned into a proven fact.
wenzelm
parents: 61664
diff changeset
   330
  The given proposition may deviate from internal meta-level equality
wenzelm
parents: 61664
diff changeset
   331
  according to the rewrite rules declared as @{attribute defn} by the
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   332
  object-logic. This usually covers object-level equality \<open>x = y\<close> and
61665
wenzelm
parents: 61664
diff changeset
   333
  equivalence \<open>A \<longleftrightarrow> B\<close>. End-users normally need not change the @{attribute
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   334
  defn} setup.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   335
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   336
  Definitions may be presented with explicit arguments on the LHS, as well as
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   337
  additional conditions, e.g.\ \<open>f x y = t\<close> instead of \<open>f \<equiv> \<lambda>x y. t\<close> and \<open>y \<noteq> 0
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   338
  \<Longrightarrow> g x y = u\<close> instead of an unrestricted \<open>g \<equiv> \<lambda>x y. u\<close>.
51585
fcd5af4aac2b added 'print_defn_rules' command;
wenzelm
parents: 51565
diff changeset
   339
61665
wenzelm
parents: 61664
diff changeset
   340
  \<^descr> \<^theory_text>\<open>print_defn_rules\<close> prints the definitional rewrite rules declared via
wenzelm
parents: 61664
diff changeset
   341
  @{attribute defn} in the current context.
51585
fcd5af4aac2b added 'print_defn_rules' command;
wenzelm
parents: 51565
diff changeset
   342
61665
wenzelm
parents: 61664
diff changeset
   343
  \<^descr> \<^theory_text>\<open>abbreviation c where eq\<close> introduces a syntactic constant which is
wenzelm
parents: 61664
diff changeset
   344
  associated with a certain term according to the meta-level equality \<open>eq\<close>.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   345
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   346
  Abbreviations participate in the usual type-inference process, but are
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   347
  expanded before the logic ever sees them. Pretty printing of terms involves
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   348
  higher-order rewriting with rules stemming from reverted abbreviations. This
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   349
  needs some care to avoid overlapping or looping syntactic replacements!
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   350
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   351
  The optional \<open>mode\<close> specification restricts output to a particular print
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   352
  mode; using ``\<open>input\<close>'' here achieves the effect of one-way abbreviations.
61665
wenzelm
parents: 61664
diff changeset
   353
  The mode may also include an ``\<^theory_text>\<open>output\<close>'' qualifier that affects the
wenzelm
parents: 61664
diff changeset
   354
  concrete syntax declared for abbreviations, cf.\ \<^theory_text>\<open>syntax\<close> in
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   355
  \secref{sec:syn-trans}.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   356
61665
wenzelm
parents: 61664
diff changeset
   357
  \<^descr> \<^theory_text>\<open>print_abbrevs\<close> prints all constant abbreviations of the current context;
wenzelm
parents: 61664
diff changeset
   358
  the ``\<open>!\<close>'' option indicates extra verbosity.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   359
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   360
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   361
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   362
section \<open>Axiomatizations \label{sec:axiomatizations}\<close>
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   363
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   364
text \<open>
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   365
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   366
    @{command_def "axiomatization"} & : & \<open>theory \<rightarrow> theory\<close> & (axiomatic!) \\
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   367
  \end{matharray}
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   368
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   369
  \<^rail>\<open>
63285
e9c777bfd78c clarified syntax;
wenzelm
parents: 63282
diff changeset
   370
    @@{command axiomatization} @{syntax vars}? (@'where' axiomatization)?
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   371
    ;
63178
b9e1d53124f5 clarified 'axiomatization';
wenzelm
parents: 63039
diff changeset
   372
    axiomatization: (@{syntax thmdecl} @{syntax prop} + @'and')
63182
b065b4833092 allow 'for' fixes for multi_specs;
wenzelm
parents: 63180
diff changeset
   373
      @{syntax spec_prems} @{syntax for_fixes}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   374
  \<close>
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   375
61665
wenzelm
parents: 61664
diff changeset
   376
  \<^descr> \<^theory_text>\<open>axiomatization c\<^sub>1 \<dots> c\<^sub>m where \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n\<close> introduces several constants
wenzelm
parents: 61664
diff changeset
   377
  simultaneously and states axiomatic properties for these. The constants are
wenzelm
parents: 61664
diff changeset
   378
  marked as being specified once and for all, which prevents additional
71567
9a29e883a934 tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents: 71166
diff changeset
   379
  specifications for the same constants later on, but it is always possible to
61665
wenzelm
parents: 61664
diff changeset
   380
  emit axiomatizations without referring to particular constants. Note that
wenzelm
parents: 61664
diff changeset
   381
  lack of precise dependency tracking of axiomatizations may disrupt the
wenzelm
parents: 61664
diff changeset
   382
  well-formedness of an otherwise definitional theory.
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   383
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   384
  Axiomatization is restricted to a global theory context: support for local
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   385
  theory targets \secref{sec:target} would introduce an extra dimension of
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   386
  uncertainty what the written specifications really are, and make it
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   387
  infeasible to argue why they are correct.
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   388
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   389
  Axiomatic specifications are required when declaring a new logical system
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   390
  within Isabelle/Pure, but in an application environment like Isabelle/HOL
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   391
  the user normally stays within definitional mechanisms provided by the logic
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   392
  and its libraries.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   393
\<close>
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   394
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
   395
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   396
section \<open>Generic declarations\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   397
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   398
text \<open>
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   399
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   400
    @{command_def "declaration"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   401
    @{command_def "syntax_declaration"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   402
    @{command_def "declare"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   403
  \end{matharray}
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   404
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   405
  Arbitrary operations on the background context may be wrapped-up as generic
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   406
  declaration elements. Since the underlying concept of local theories may be
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   407
  subject to later re-interpretation, there is an additional dependency on a
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   408
  morphism that tells the difference of the original declaration context wrt.\
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   409
  the application context encountered later on. A fact declaration is an
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   410
  important special case: it consists of a theorem which is applied to the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   411
  context by means of an attribute.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   412
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   413
  \<^rail>\<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
   414
    (@@{command declaration} | @@{command syntax_declaration})
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 59781
diff changeset
   415
      ('(' @'pervasive' ')')? \<newline> @{syntax text}
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   416
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   417
    @@{command declare} (@{syntax thms} + @'and')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   418
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   419
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   420
  \<^descr> \<^theory_text>\<open>declaration d\<close> adds the declaration function \<open>d\<close> of ML type \<^ML_type>\<open>declaration\<close>, to the current local theory under construction. In later
61665
wenzelm
parents: 61664
diff changeset
   421
  application contexts, the function is transformed according to the morphisms
wenzelm
parents: 61664
diff changeset
   422
  being involved in the interpretation hierarchy.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   423
61665
wenzelm
parents: 61664
diff changeset
   424
  If the \<^theory_text>\<open>(pervasive)\<close> option is given, the corresponding declaration is
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   425
  applied to all possible contexts involved, including the global background
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   426
  theory.
33516
0855a09bc5cf added "declaration (pervasive)";
wenzelm
parents: 31914
diff changeset
   427
61665
wenzelm
parents: 61664
diff changeset
   428
  \<^descr> \<^theory_text>\<open>syntax_declaration\<close> is similar to \<^theory_text>\<open>declaration\<close>, but is meant to affect
wenzelm
parents: 61664
diff changeset
   429
  only ``syntactic'' tools by convention (such as notation and type-checking
wenzelm
parents: 61664
diff changeset
   430
  information).
40784
177e8cea3e09 added 'syntax_declaration' command;
wenzelm
parents: 40256
diff changeset
   431
61665
wenzelm
parents: 61664
diff changeset
   432
  \<^descr> \<^theory_text>\<open>declare thms\<close> declares theorems to the current local theory context. No
wenzelm
parents: 61664
diff changeset
   433
  theorem binding is involved here, unlike \<^theory_text>\<open>lemmas\<close> (cf.\
wenzelm
parents: 61664
diff changeset
   434
  \secref{sec:theorems}), so \<^theory_text>\<open>declare\<close> only has the effect of applying
wenzelm
parents: 61664
diff changeset
   435
  attributes as included in the theorem specification.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   436
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   437
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   438
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   439
section \<open>Locales \label{sec:locale}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   440
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   441
text \<open>
54049
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   442
  A locale is a functor that maps parameters (including implicit type
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   443
  parameters) and a specification to a list of declarations. The syntax of
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   444
  locales is modeled after the Isar proof context commands (cf.\
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   445
  \secref{sec:proof-context}).
54049
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   446
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   447
  Locale hierarchies are supported by maintaining a graph of dependencies
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   448
  between locale instances in the global theory. Dependencies may be
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   449
  introduced through import (where a locale is defined as sublocale of the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   450
  imported instances) or by proving that an existing locale is a sublocale of
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   451
  one or several locale instances.
54049
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   452
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   453
  A locale may be opened with the purpose of appending to its list of
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   454
  declarations (cf.\ \secref{sec:target}). When opening a locale declarations
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   455
  from all dependencies are collected and are presented as a local theory. In
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   456
  this process, which is called \<^emph>\<open>roundup\<close>, redundant locale instances are
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   457
  omitted. A locale instance is redundant if it is subsumed by an instance
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   458
  encountered earlier. A more detailed description of this process is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   459
  available elsewhere @{cite Ballarin2014}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   460
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   461
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   462
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   463
subsection \<open>Locale expressions \label{sec:locale-expr}\<close>
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   464
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   465
text \<open>
61606
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   466
  A \<^emph>\<open>locale expression\<close> denotes a context composed of instances of existing
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   467
  locales. The context consists of the declaration elements from the locale
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   468
  instances. Redundant locale instances are omitted according to roundup.
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   469
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   470
  \<^rail>\<open>
59785
4e6ab5831cc0 clarified syntax category "fixes";
wenzelm
parents: 59783
diff changeset
   471
    @{syntax_def locale_expr}: (instance + '+') @{syntax for_fixes}
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   472
    ;
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   473
    instance: (qualifier ':')? @{syntax name} (pos_insts | named_insts) \<newline>
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   474
      rewrites?
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   475
    ;
61606
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   476
    qualifier: @{syntax name} ('?')?
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   477
    ;
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   478
    pos_insts: ('_' | @{syntax term})*
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   479
    ;
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   480
    named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and')
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   481
    ;
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   482
    rewrites: @'rewrites' (@{syntax thmdecl}? @{syntax prop} + @'and')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   483
  \<close>
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   484
61606
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   485
  A locale instance consists of a reference to a locale and either positional
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   486
  or named parameter instantiations optionally followed by rewrites clauses.
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   487
  Identical instantiations (that is, those
61606
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   488
  that instantiate a parameter by itself) may be omitted. The notation ``\<open>_\<close>''
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   489
  enables to omit the instantiation for a parameter inside a positional
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   490
  instantiation.
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   491
61606
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   492
  Terms in instantiations are from the context the locale expressions is
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   493
  declared in. Local names may be added to this context with the optional
61665
wenzelm
parents: 61664
diff changeset
   494
  \<^theory_text>\<open>for\<close> clause. This is useful for shadowing names bound in outer contexts,
wenzelm
parents: 61664
diff changeset
   495
  and for declaring syntax. In addition, syntax declarations from one instance
wenzelm
parents: 61664
diff changeset
   496
  are effective when parsing subsequent instances of the same expression.
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   497
61606
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   498
  Instances have an optional qualifier which applies to names in declarations.
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   499
  Names include local definitions and theorem names. If present, the qualifier
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   500
  itself is either mandatory (default) or non-mandatory (when followed by
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   501
  ``\<^verbatim>\<open>?\<close>''). Non-mandatory means that the qualifier may be omitted on input.
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   502
  Qualifiers only affect name spaces; they play no role in determining whether
6d5213bd9709 uniform mandatory qualifier for all locale expressions, including 'statespace' parent;
wenzelm
parents: 61566
diff changeset
   503
  one locale instance subsumes another.
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   504
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   505
  Rewrite clauses amend instances with equations that act as rewrite rules.
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   506
  This is particularly useful for changing concepts introduced through
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   507
  definitions. Rewrite clauses are available only in interpretation commands
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   508
  (see \secref{sec:locale-interpretation} below) and must be proved the user.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   509
\<close>
33846
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   510
cff41e82e3df Updated locale documentation.
ballarin
parents: 33516
diff changeset
   511
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   512
subsection \<open>Locale declarations\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   513
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   514
text \<open>
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   515
  \begin{tabular}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   516
    @{command_def "locale"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   517
    @{command_def "experiment"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   518
    @{command_def "print_locale"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   519
    @{command_def "print_locales"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   520
    @{command_def "locale_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
70608
d997c7ba3305 Tracing of locale activation.
ballarin
parents: 70560
diff changeset
   521
  \end{tabular}
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   522
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   523
  \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes}
28787
8ea7403147c5 removed "includes" element (lost update?);
wenzelm
parents: 28768
diff changeset
   524
  \indexisarelem{defines}\indexisarelem{notes}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   525
  \<^rail>\<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
   526
    @@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'?
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   527
    ;
59901
840d03805755 added command 'experiment';
wenzelm
parents: 59785
diff changeset
   528
    @@{command experiment} (@{syntax context_elem}*) @'begin'
840d03805755 added command 'experiment';
wenzelm
parents: 59785
diff changeset
   529
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   530
    @@{command print_locale} '!'? @{syntax name}
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   531
    ;
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   532
    @@{command print_locales} ('!'?)
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   533
    ;
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   534
    @{syntax_def locale}: @{syntax context_elem}+ |
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   535
      @{syntax_ref "opening"} ('+' (@{syntax context_elem}+))? |
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   536
      @{syntax locale_expr} @{syntax_ref "opening"}? ('+' (@{syntax context_elem}+))?
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   537
    ;
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   538
    @{syntax_def context_elem}:
63285
e9c777bfd78c clarified syntax;
wenzelm
parents: 63282
diff changeset
   539
      @'fixes' @{syntax vars} |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
   540
      @'constrains' (@{syntax name} '::' @{syntax type} + @'and') |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
   541
      @'assumes' (@{syntax props} + @'and') |
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
   542
      @'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') |
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   543
      @'notes' (@{syntax thmdef}? @{syntax thms} + @'and')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   544
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   545
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   546
  \<^descr> \<^theory_text>\<open>locale loc = import opening bundles + body\<close> defines a new locale \<open>loc\<close>
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   547
  as a context consisting of a certain view of existing locales (\<open>import\<close>) plus some
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   548
  additional elements (\<open>body\<close>) with declaration \<open>bundles\<close> enriching the context
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   549
  of the command itself. All \<open>import\<close>, \<open>bundles\<close> and \<open>body\<close> are optional; the
61665
wenzelm
parents: 61664
diff changeset
   550
  degenerate form \<^theory_text>\<open>locale loc\<close> defines an empty locale, which may still be
wenzelm
parents: 61664
diff changeset
   551
  useful to collect declarations of facts later on. Type-inference on locale
wenzelm
parents: 61664
diff changeset
   552
  expressions automatically takes care of the most general typing that the
wenzelm
parents: 61664
diff changeset
   553
  combined context elements may acquire.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   554
71567
9a29e883a934 tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents: 71166
diff changeset
   555
  The \<open>import\<close> consists of a locale expression; see \secref{sec:locale-expr}
61665
wenzelm
parents: 61664
diff changeset
   556
  above. Its \<^theory_text>\<open>for\<close> clause defines the parameters of \<open>import\<close>. These are
wenzelm
parents: 61664
diff changeset
   557
  parameters of the defined locale. Locale parameters whose instantiation is
wenzelm
parents: 61664
diff changeset
   558
  omitted automatically extend the (possibly empty) \<^theory_text>\<open>for\<close> clause: they are
wenzelm
parents: 61664
diff changeset
   559
  inserted at its beginning. This means that these parameters may be referred
wenzelm
parents: 61664
diff changeset
   560
  to from within the expression and also in the subsequent context elements
wenzelm
parents: 61664
diff changeset
   561
  and provides a notational convenience for the inheritance of parameters in
wenzelm
parents: 61664
diff changeset
   562
  locale declarations.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   563
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   564
  Declarations from \<open>bundles\<close>, see \secref{sec:bundle}, are effective in the
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   565
  entire command including a subsequent \<^theory_text>\<open>begin\<close> / \<^theory_text>\<open>end\<close> block, but they do
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   566
  not contribute to the declarations stored in the locale.
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   567
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   568
  The \<open>body\<close> consists of context elements:
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   569
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   570
    \<^descr> @{element "fixes"}~\<open>x :: \<tau> (mx)\<close> declares a local parameter of type \<open>\<tau>\<close>
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   571
    and mixfix annotation \<open>mx\<close> (both are optional). The special syntax
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   572
    declaration ``\<open>(\<close>@{keyword_ref "structure"}\<open>)\<close>'' means that \<open>x\<close> may be
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   573
    referenced implicitly in this context.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   574
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   575
    \<^descr> @{element "constrains"}~\<open>x :: \<tau>\<close> introduces a type constraint \<open>\<tau>\<close> on the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   576
    local parameter \<open>x\<close>. This element is deprecated. The type constraint
61665
wenzelm
parents: 61664
diff changeset
   577
    should be introduced in the \<^theory_text>\<open>for\<close> clause or the relevant @{element
wenzelm
parents: 61664
diff changeset
   578
    "fixes"} element.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   579
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   580
    \<^descr> @{element "assumes"}~\<open>a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n\<close> introduces local premises, similar
61665
wenzelm
parents: 61664
diff changeset
   581
    to \<^theory_text>\<open>assume\<close> within a proof (cf.\ \secref{sec:proof-context}).
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   582
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   583
    \<^descr> @{element "defines"}~\<open>a: x \<equiv> t\<close> defines a previously declared parameter.
63039
1a20fd9cf281 added Isar command 'define';
wenzelm
parents: 62969
diff changeset
   584
    This is similar to \<^theory_text>\<open>define\<close> within a proof (cf.\
1a20fd9cf281 added Isar command 'define';
wenzelm
parents: 62969
diff changeset
   585
    \secref{sec:proof-context}), but @{element "defines"} is restricted to
1a20fd9cf281 added Isar command 'define';
wenzelm
parents: 62969
diff changeset
   586
    Pure equalities and the defined variable needs to be declared beforehand
1a20fd9cf281 added Isar command 'define';
wenzelm
parents: 62969
diff changeset
   587
    via @{element "fixes"}. The left-hand side of the equation may have
1a20fd9cf281 added Isar command 'define';
wenzelm
parents: 62969
diff changeset
   588
    additional arguments, e.g.\ ``@{element "defines"}~\<open>f x\<^sub>1 \<dots> x\<^sub>n \<equiv> t\<close>'',
1a20fd9cf281 added Isar command 'define';
wenzelm
parents: 62969
diff changeset
   589
    which need to be free in the context.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   590
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   591
    \<^descr> @{element "notes"}~\<open>a = b\<^sub>1 \<dots> b\<^sub>n\<close> reconsiders facts within a local
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   592
    context. Most notably, this may include arbitrary declarations in any
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   593
    attribute specifications included here, e.g.\ a local @{attribute simp}
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   594
    rule.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   595
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   596
  Both @{element "assumes"} and @{element "defines"} elements contribute to
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   597
  the locale specification. When defining an operation derived from the
61665
wenzelm
parents: 61664
diff changeset
   598
  parameters, \<^theory_text>\<open>definition\<close> (\secref{sec:term-definitions}) is usually more
wenzelm
parents: 61664
diff changeset
   599
  appropriate.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   600
61665
wenzelm
parents: 61664
diff changeset
   601
  Note that ``\<^theory_text>\<open>(is p\<^sub>1 \<dots> p\<^sub>n)\<close>'' patterns given in the syntax of @{element
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   602
  "assumes"} and @{element "defines"} above are illegal in locale definitions.
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   603
  In the long goal format of \secref{sec:goals}, term bindings may be included
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   604
  as expected, though.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   605
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   606
  \<^medskip>
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   607
  Locale specifications are ``closed up'' by turning the given text into a
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   608
  predicate definition \<open>loc_axioms\<close> and deriving the original assumptions as
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   609
  local lemmas (modulo local definitions). The predicate statement covers only
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   610
  the newly specified assumptions, omitting the content of included locale
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   611
  expressions. The full cumulative view is only provided on export, involving
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   612
  another predicate \<open>loc\<close> that refers to the complete specification text.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
   613
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   614
  In any case, the predicate arguments are those locale parameters that
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   615
  actually occur in the respective piece of text. Also these predicates
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   616
  operate at the meta-level in theory, but the locale packages attempts to
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   617
  internalize statements according to the object-logic setup (e.g.\ replacing
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   618
  \<open>\<And>\<close> by \<open>\<forall>\<close>, and \<open>\<Longrightarrow>\<close> by \<open>\<longrightarrow>\<close> in HOL; see also \secref{sec:object-logic}).
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   619
  Separate introduction rules \<open>loc_axioms.intro\<close> and \<open>loc.intro\<close> are provided
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   620
  as well.
59901
840d03805755 added command 'experiment';
wenzelm
parents: 59785
diff changeset
   621
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   622
  \<^descr> \<^theory_text>\<open>experiment body begin\<close> opens an anonymous locale context with private
61665
wenzelm
parents: 61664
diff changeset
   623
  naming policy. Specifications in its body are inaccessible from outside.
wenzelm
parents: 61664
diff changeset
   624
  This is useful to perform experiments, without polluting the name space.
59901
840d03805755 added command 'experiment';
wenzelm
parents: 59785
diff changeset
   625
69018
wenzelm
parents: 68278
diff changeset
   626
  \<^descr> \<^theory_text>\<open>print_locale "locale"\<close> prints the contents of the named locale. The
61665
wenzelm
parents: 61664
diff changeset
   627
  command omits @{element "notes"} elements by default. Use \<^theory_text>\<open>print_locale!\<close>
wenzelm
parents: 61664
diff changeset
   628
  to have them included.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   629
61665
wenzelm
parents: 61664
diff changeset
   630
  \<^descr> \<^theory_text>\<open>print_locales\<close> prints the names of all locales of the current theory;
wenzelm
parents: 61664
diff changeset
   631
  the ``\<open>!\<close>'' option indicates extra verbosity.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   632
61665
wenzelm
parents: 61664
diff changeset
   633
  \<^descr> \<^theory_text>\<open>locale_deps\<close> visualizes all locales and their relations as a Hasse
wenzelm
parents: 61664
diff changeset
   634
  diagram. This includes locales defined as type classes (\secref{sec:class}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   635
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   636
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   637
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   638
subsection \<open>Locale interpretation \label{sec:locale-interpretation}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   639
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   640
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   641
  \begin{matharray}{rcl}
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   642
    @{command "interpretation"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   643
    @{command_def "interpret"} & : & \<open>proof(state) | proof(chain) \<rightarrow> proof(prove)\<close> \\
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   644
    @{command_def "global_interpretation"} & : & \<open>theory | local_theory \<rightarrow> proof(prove)\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   645
    @{command_def "sublocale"} & : & \<open>theory | local_theory \<rightarrow> proof(prove)\<close> \\
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   646
    @{command_def "print_interps"}\<open>\<^sup>*\<close> & :  & \<open>context \<rightarrow>\<close> \\
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   647
    @{method_def intro_locales} & : & \<open>method\<close> \\
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   648
    @{method_def unfold_locales} & : & \<open>method\<close> \\
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   649
    @{attribute_def trace_locales} & : & \mbox{\<open>attribute\<close> \quad default \<open>false\<close>} \\
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   650
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   651
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   652
  Locales may be instantiated, and the resulting instantiated declarations
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   653
  added to the current context. This requires proof (of the instantiated
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   654
  specification) and is called \<^emph>\<open>locale interpretation\<close>. Interpretation is
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   655
  possible within arbitrary local theories (\<^theory_text>\<open>interpretation\<close>), within proof
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   656
  bodies (\<^theory_text>\<open>interpret\<close>), into global theories (\<^theory_text>\<open>global_interpretation\<close>) and
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   657
  into locales (\<^theory_text>\<open>sublocale\<close>).
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   658
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   659
  \<^rail>\<open>
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   660
    @@{command interpretation} @{syntax locale_expr}
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   661
    ;
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   662
    @@{command interpret} @{syntax locale_expr}
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   663
    ;
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   664
    @@{command global_interpretation} @{syntax locale_expr} definitions?
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   665
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   666
    @@{command sublocale} (@{syntax name} ('<' | '\<subseteq>'))? @{syntax locale_expr} \<newline>
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   667
      definitions?
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   668
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   669
    @@{command print_interps} @{syntax name}
41434
710cdb9e0d17 Documentation for 'interpret' and 'sublocale' with mixins.
ballarin
parents: 41249
diff changeset
   670
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
   671
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   672
    definitions: @'defines' (@{syntax thmdecl}? @{syntax name} \<newline>
70058
wenzelm
parents: 70057
diff changeset
   673
      @{syntax mixfix}? '=' @{syntax term} + @'and');
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   674
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   675
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   676
  The core of each interpretation command is a locale expression \<open>expr\<close>; the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   677
  command generates proof obligations for the instantiated specifications.
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   678
  Once these are discharged by the user, instantiated declarations (in
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   679
  particular, facts) are added to the context in a post-processing phase, in a
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   680
  manner specific to each command.
53368
92b9965f479d Clarifies documentation of interpretation in local theories.
ballarin
parents: 53015
diff changeset
   681
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   682
  Interpretation commands are aware of interpretations that are already
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   683
  active: post-processing is achieved through a variant of roundup that takes
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   684
  interpretations of the current global or local theory into account. In order
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   685
  to simplify the proof obligations according to existing interpretations use
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   686
  methods @{method intro_locales} or @{method unfold_locales}.
53368
92b9965f479d Clarifies documentation of interpretation in local theories.
ballarin
parents: 53015
diff changeset
   687
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   688
  Rewrites clauses \<^theory_text>\<open>rewrites eqns\<close> occur within expressions. They amend the
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   689
  morphism through which a locale instance is interpreted with rewrite rules,
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   690
  also called rewrite morphisms. This is particularly useful for interpreting
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   691
  concepts introduced through definitions. The equations must be proved the
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   692
  user. To enable syntax of the instantiated locale within the equation, while
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   693
  reading a locale expression, equations of a locale instance are read in a
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   694
  temporary context where the instance is already activated. If activation
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   695
  fails, typically due to duplicate constant declarations, processing falls
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   696
  back to reading the equation first.
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   697
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   698
  Given definitions \<open>defs\<close> produce corresponding definitions in the local
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   699
  theory's underlying target \<^emph>\<open>and\<close> amend the morphism with rewrite rules
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   700
  stemming from the symmetric of those definitions. Hence these need not be
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   701
  proved explicitly the user. Such rewrite definitions are a even more useful
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   702
  device for interpreting concepts introduced through definitions, but they
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   703
  are only supported for interpretation commands operating in a local theory
61823
5daa82ba4078 clarified terminology
haftmann
parents: 61706
diff changeset
   704
  whose implementing target actually supports this.  Note that despite
5daa82ba4078 clarified terminology
haftmann
parents: 61706
diff changeset
   705
  the suggestive \<^theory_text>\<open>and\<close> connective, \<open>defs\<close>
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   706
  are processed sequentially without mutual recursion.
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   707
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   708
  \<^descr> \<^theory_text>\<open>interpretation expr\<close> interprets \<open>expr\<close> into a local theory
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   709
  such that its lifetime is limited to the current context block (e.g. a
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   710
  locale or unnamed context). At the closing @{command end} of the block the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   711
  interpretation and its declarations disappear. Hence facts based on
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   712
  interpretation can be established without creating permanent links to the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   713
  interpreted locale instances, as would be the case with @{command
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   714
  sublocale}.
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   715
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   716
  When used on the level of a global theory, there is no end of a current
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   717
  context block, hence \<^theory_text>\<open>interpretation\<close> behaves identically to
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   718
  \<^theory_text>\<open>global_interpretation\<close> then.
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   719
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   720
  \<^descr> \<^theory_text>\<open>interpret expr\<close> interprets \<open>expr\<close> into a proof context:
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   721
  the interpretation and its declarations disappear when closing the current
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   722
  proof block. Note that for \<^theory_text>\<open>interpret\<close> the \<open>eqns\<close> should be explicitly
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   723
  universally quantified.
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   724
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   725
  \<^descr> \<^theory_text>\<open>global_interpretation expr defines defs\<close> interprets \<open>expr\<close>
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   726
  into a global theory.
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   727
53369
b4bcac7d065b Clarifies that interpretation does not only apply to facts, but to declaratoins in general.
ballarin
parents: 53368
diff changeset
   728
  When adding declarations to locales, interpreted versions of these
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   729
  declarations are added to the global theory for all interpretations in the
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61853
diff changeset
   730
  global theory as well. That is, interpretations into global theories
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   731
  dynamically participate in any declarations added to locales.
54049
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   732
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   733
  Free variables in the interpreted expression are allowed. They are turned
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   734
  into schematic variables in the generated declarations. In order to use a
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   735
  free variable whose name is already bound in the context --- for example,
61665
wenzelm
parents: 61664
diff changeset
   736
  because a constant of that name exists --- add it to the \<^theory_text>\<open>for\<close> clause.
48824
45d0e40b07af Clarification: free variables allowed in interpreted locale instances.
ballarin
parents: 47484
diff changeset
   737
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   738
  \<^descr> \<^theory_text>\<open>sublocale name \<subseteq> expr defines defs\<close> interprets \<open>expr\<close>
61706
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   739
  into the locale \<open>name\<close>. A proof that the specification of \<open>name\<close> implies the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   740
  specification of \<open>expr\<close> is required. As in the localized version of the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   741
  theorem command, the proof is in the context of \<open>name\<close>. After the proof
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   742
  obligation has been discharged, the locale hierarchy is changed as if \<open>name\<close>
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   743
  imported \<open>expr\<close> (hence the name \<^theory_text>\<open>sublocale\<close>). When the context of \<open>name\<close> is
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   744
  subsequently entered, traversing the locale hierarchy will involve the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   745
  locale instances of \<open>expr\<close>, and their declarations will be added to the
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   746
  context. This makes \<^theory_text>\<open>sublocale\<close> dynamic: extensions of a locale that is
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   747
  instantiated in \<open>expr\<close> may take place after the \<^theory_text>\<open>sublocale\<close> declaration and
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   748
  still become available in the context. Circular \<^theory_text>\<open>sublocale\<close> declarations
a1510dfb9bfa tuned whitespace;
wenzelm
parents: 61673
diff changeset
   749
  are allowed as long as they do not lead to infinite chains.
41434
710cdb9e0d17 Documentation for 'interpret' and 'sublocale' with mixins.
ballarin
parents: 41249
diff changeset
   750
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   751
  If interpretations of \<open>name\<close> exist in the current global theory, the command
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   752
  adds interpretations for \<open>expr\<close> as well, with the same qualifier, although
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   753
  only for fragments of \<open>expr\<close> that are not interpreted in the theory already.
41434
710cdb9e0d17 Documentation for 'interpret' and 'sublocale' with mixins.
ballarin
parents: 41249
diff changeset
   754
67764
0f8cb5568b63 Drop rewrites after defines in interpretations.
ballarin
parents: 67740
diff changeset
   755
  Rewrites clauses in the expression or rewrite definitions \<open>defs\<close> can help break
67740
b6ce18784872 Proper rewrite morphisms in locale instances.
ballarin
parents: 67722
diff changeset
   756
  infinite chains induced by circular \<^theory_text>\<open>sublocale\<close> declarations.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   757
61665
wenzelm
parents: 61664
diff changeset
   758
  In a named context block the \<^theory_text>\<open>sublocale\<close> command may also be used, but the
wenzelm
parents: 61664
diff changeset
   759
  locale argument must be omitted. The command then refers to the locale (or
wenzelm
parents: 61664
diff changeset
   760
  class) target of the context block.
51747
e4b5bebe5235 documentation and NEWS
haftmann
parents: 51657
diff changeset
   761
71166
c9433e8e314e Remove diagnostic command 'print_dependencies'.
ballarin
parents: 70608
diff changeset
   762
  \<^descr> \<^theory_text>\<open>print_interps name\<close> lists all interpretations of locale \<open>name\<close> in the
61665
wenzelm
parents: 61664
diff changeset
   763
  current theory or proof context, including those due to a combination of an
wenzelm
parents: 61664
diff changeset
   764
  \<^theory_text>\<open>interpretation\<close> or \<^theory_text>\<open>interpret\<close> and one or several \<^theory_text>\<open>sublocale\<close>
wenzelm
parents: 61664
diff changeset
   765
  declarations.
33867
52643d0f856d Typos in documenation.
ballarin
parents: 33846
diff changeset
   766
72804
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   767
  \<^descr> @{method intro_locales} and @{method unfold_locales} repeatedly expand all
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   768
  introduction rules of locale predicates of the theory. While @{method
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   769
  intro_locales} only applies the \<open>loc.intro\<close> introduction rules and therefore
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   770
  does not descend to assumptions, @{method unfold_locales} is more aggressive
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   771
  and applies \<open>loc_axioms.intro\<close> as well. Both methods are aware of locale
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   772
  specifications entailed by the context, both from target statements, and
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   773
  from interpretations (see below). New goals that are entailed by the current
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   774
  context are discharged automatically.
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   775
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   776
  While @{method unfold_locales} is part of the default method for \<^theory_text>\<open>proof\<close>
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   777
  and often invoked ``behind the scenes'', @{method intro_locales} helps
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   778
  understand which proof obligations originated from which locale instances.
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   779
  The latter method is useful while developing proofs but rare in finished
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   780
  developments.
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   781
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   782
  \<^descr> @{attribute trace_locales}, when set to \<open>true\<close>, prints the locale
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   783
  instances activated during roundup. Use this when locale commands yield
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   784
  obscure errors or for understanding local theories created by complex locale
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   785
  hierarchies.
d7393c35aa5d Highlight that 'intro_locales' and 'trace_locales' are useful debug features, plus minor clarifications.
ballarin
parents: 72739
diff changeset
   786
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   787
  \begin{warn}
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   788
    If a global theory inherits declarations (body elements) for a locale from
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   789
    one parent and an interpretation of that locale from another parent, the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   790
    interpretation will not be applied to the declarations.
54049
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   791
  \end{warn}
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   792
566b769c3477 Streamlined locales reference material.
ballarin
parents: 53536
diff changeset
   793
  \begin{warn}
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   794
    Since attributes are applied to interpreted theorems, interpretation may
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   795
    modify the context of common proof tools, e.g.\ the Simplifier or
61665
wenzelm
parents: 61664
diff changeset
   796
    Classical Reasoner. As the behaviour of such tools is \<^emph>\<open>not\<close> stable under
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   797
    interpretation morphisms, manual declarations might have to be added to
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   798
    the target context of the interpretation to revert such declarations.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   799
  \end{warn}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   800
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   801
  \begin{warn}
51747
e4b5bebe5235 documentation and NEWS
haftmann
parents: 51657
diff changeset
   802
    An interpretation in a local theory or proof context may subsume previous
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   803
    interpretations. This happens if the same specification fragment is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   804
    interpreted twice and the instantiation of the second interpretation is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   805
    more general than the interpretation of the first. The locale package does
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   806
    not attempt to remove subsumed interpretations.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   807
  \end{warn}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   808
61673
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   809
  \begin{warn}
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   810
    While \<^theory_text>\<open>interpretation (in c) \<dots>\<close> is admissible, it is not useful since
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   811
    its result is discarded immediately.
fd4ac1530d63 represent both algebraic and local-theory views on locale interpretation in interfaces
haftmann
parents: 61671
diff changeset
   812
  \end{warn}
59003
16d92d37a8a1 documentation stubs about permanent_interpretation
haftmann
parents: 58724
diff changeset
   813
\<close>
16d92d37a8a1 documentation stubs about permanent_interpretation
haftmann
parents: 58724
diff changeset
   814
16d92d37a8a1 documentation stubs about permanent_interpretation
haftmann
parents: 58724
diff changeset
   815
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   816
section \<open>Classes \label{sec:class}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   817
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   818
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   819
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   820
    @{command_def "class"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   821
    @{command_def "instantiation"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   822
    @{command_def "instance"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   823
    @{command "instance"} & : & \<open>theory \<rightarrow> proof(prove)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   824
    @{command_def "subclass"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   825
    @{command_def "print_classes"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   826
    @{command_def "class_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   827
    @{method_def intro_classes} & : & \<open>method\<close> \\
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   828
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   829
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   830
  A class is a particular locale with \<^emph>\<open>exactly one\<close> type variable \<open>\<alpha>\<close>. Beyond
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   831
  the underlying locale, a corresponding type class is established which is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   832
  interpreted logically as axiomatic type class @{cite "Wenzel:1997:TPHOL"}
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   833
  whose logical content are the assumptions of the locale. Thus, classes
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   834
  provide the full generality of locales combined with the commodity of type
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   835
  classes (notably type-inference). See @{cite "isabelle-classes"} for a short
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   836
  tutorial.
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
   837
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   838
  \<^rail>\<open>
42704
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42651
diff changeset
   839
    @@{command class} class_spec @'begin'?
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42651
diff changeset
   840
    ;
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42651
diff changeset
   841
    class_spec: @{syntax name} '='
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   842
      ((@{syntax name} @{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) |
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   843
        @{syntax name} @{syntax_ref "opening"}? |
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   844
        @{syntax_ref "opening"}? '+' (@{syntax context_elem}+))
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   845
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   846
    @@{command instantiation} (@{syntax name} + @'and') '::' @{syntax arity} @'begin'
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   847
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   848
    @@{command instance} (() | (@{syntax name} + @'and') '::' @{syntax arity} |
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   849
      @{syntax name} ('<' | '\<subseteq>') @{syntax name} )
31681
127e8a8b8cde refined section concerning classes
haftmann
parents: 31047
diff changeset
   850
    ;
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
   851
    @@{command subclass} @{syntax name}
58202
be1d10595b7b restrictive options for class dependencies
haftmann
parents: 57947
diff changeset
   852
    ;
60091
9feddd64183e tuned signature;
wenzelm
parents: 60090
diff changeset
   853
    @@{command class_deps} (class_bounds class_bounds?)?
60090
75ec8fd5d2bf misc tuning and clarification;
wenzelm
parents: 59990
diff changeset
   854
    ;
60091
9feddd64183e tuned signature;
wenzelm
parents: 60090
diff changeset
   855
    class_bounds: @{syntax sort} | '(' (@{syntax sort} + @'|') ')'
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
   856
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   857
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   858
  \<^descr> \<^theory_text>\<open>class c = superclasses bundles + body\<close> defines a new class \<open>c\<close>, inheriting from
61665
wenzelm
parents: 61664
diff changeset
   859
  \<open>superclasses\<close>. This introduces a locale \<open>c\<close> with import of all locales
wenzelm
parents: 61664
diff changeset
   860
  \<open>superclasses\<close>.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   861
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   862
  Any @{element "fixes"} in \<open>body\<close> are lifted to the global theory level
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   863
  (\<^emph>\<open>class operations\<close> \<open>f\<^sub>1, \<dots>, f\<^sub>n\<close> of class \<open>c\<close>), mapping the local type
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   864
  parameter \<open>\<alpha>\<close> to a schematic type variable \<open>?\<alpha> :: c\<close>.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   865
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   866
  Likewise, @{element "assumes"} in \<open>body\<close> are also lifted, mapping each local
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   867
  parameter \<open>f :: \<tau>[\<alpha>]\<close> to its corresponding global constant \<open>f :: \<tau>[?\<alpha> ::
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   868
  c]\<close>. The corresponding introduction rule is provided as
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   869
  \<open>c_class_axioms.intro\<close>. This rule should be rarely needed directly --- the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   870
  @{method intro_classes} method takes care of the details of class membership
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   871
  proofs.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   872
72739
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   873
  Optionally given \<open>bundles\<close> take effect in the surface context within the
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   874
  \<open>body\<close> and the potentially following \<^theory_text>\<open>begin\<close> / \<^theory_text>\<open>end\<close> block.
e7c2848b78e8 refined syntax for bundle mixins for locale and class specifications
haftmann
parents: 71926
diff changeset
   875
61665
wenzelm
parents: 61664
diff changeset
   876
  \<^descr> \<^theory_text>\<open>instantiation t :: (s\<^sub>1, \<dots>, s\<^sub>n)s begin\<close> opens a target (cf.\
wenzelm
parents: 61664
diff changeset
   877
  \secref{sec:target}) which allows to specify class operations \<open>f\<^sub>1, \<dots>, f\<^sub>n\<close>
wenzelm
parents: 61664
diff changeset
   878
  corresponding to sort \<open>s\<close> at the particular type instance \<open>(\<alpha>\<^sub>1 :: s\<^sub>1, \<dots>,
wenzelm
parents: 61664
diff changeset
   879
  \<alpha>\<^sub>n :: s\<^sub>n) t\<close>. A plain \<^theory_text>\<open>instance\<close> command in the target body poses a goal
wenzelm
parents: 61664
diff changeset
   880
  stating these type arities. The target is concluded by an @{command_ref
wenzelm
parents: 61664
diff changeset
   881
  (local) "end"} command.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   882
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   883
  Note that a list of simultaneous type constructors may be given; this
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   884
  corresponds nicely to mutually recursive type definitions, e.g.\ in
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   885
  Isabelle/HOL.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   886
61665
wenzelm
parents: 61664
diff changeset
   887
  \<^descr> \<^theory_text>\<open>instance\<close> in an instantiation target body sets up a goal stating the
wenzelm
parents: 61664
diff changeset
   888
  type arities claimed at the opening \<^theory_text>\<open>instantiation\<close>. The proof would
wenzelm
parents: 61664
diff changeset
   889
  usually proceed by @{method intro_classes}, and then establish the
wenzelm
parents: 61664
diff changeset
   890
  characteristic theorems of the type classes involved. After finishing the
wenzelm
parents: 61664
diff changeset
   891
  proof, the background theory will be augmented by the proven type arities.
31681
127e8a8b8cde refined section concerning classes
haftmann
parents: 31047
diff changeset
   892
61665
wenzelm
parents: 61664
diff changeset
   893
  On the theory level, \<^theory_text>\<open>instance t :: (s\<^sub>1, \<dots>, s\<^sub>n)s\<close> provides a convenient
wenzelm
parents: 61664
diff changeset
   894
  way to instantiate a type class with no need to specify operations: one can
wenzelm
parents: 61664
diff changeset
   895
  continue with the instantiation proof immediately.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   896
61665
wenzelm
parents: 61664
diff changeset
   897
  \<^descr> \<^theory_text>\<open>subclass c\<close> in a class context for class \<open>d\<close> sets up a goal stating that
wenzelm
parents: 61664
diff changeset
   898
  class \<open>c\<close> is logically contained in class \<open>d\<close>. After finishing the proof,
wenzelm
parents: 61664
diff changeset
   899
  class \<open>d\<close> is proven to be subclass \<open>c\<close> and the locale \<open>c\<close> is interpreted
wenzelm
parents: 61664
diff changeset
   900
  into \<open>d\<close> simultaneously.
31681
127e8a8b8cde refined section concerning classes
haftmann
parents: 31047
diff changeset
   901
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   902
  A weakened form of this is available through a further variant of @{command
61665
wenzelm
parents: 61664
diff changeset
   903
  instance}: \<^theory_text>\<open>instance c\<^sub>1 \<subseteq> c\<^sub>2\<close> opens a proof that class \<open>c\<^sub>2\<close> implies
wenzelm
parents: 61664
diff changeset
   904
  \<open>c\<^sub>1\<close> without reference to the underlying locales; this is useful if the
wenzelm
parents: 61664
diff changeset
   905
  properties to prove the logical connection are not sufficient on the locale
wenzelm
parents: 61664
diff changeset
   906
  level but on the theory level.
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   907
61665
wenzelm
parents: 61664
diff changeset
   908
  \<^descr> \<^theory_text>\<open>print_classes\<close> prints all classes in the current theory.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   909
61665
wenzelm
parents: 61664
diff changeset
   910
  \<^descr> \<^theory_text>\<open>class_deps\<close> visualizes classes and their subclass relations as a
wenzelm
parents: 61664
diff changeset
   911
  directed acyclic graph. By default, all classes from the current theory
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   912
  context are show. This may be restricted by optional bounds as follows:
61665
wenzelm
parents: 61664
diff changeset
   913
  \<^theory_text>\<open>class_deps upper\<close> or \<^theory_text>\<open>class_deps upper lower\<close>. A class is visualized, iff
wenzelm
parents: 61664
diff changeset
   914
  it is a subclass of some sort from \<open>upper\<close> and a superclass of some sort
wenzelm
parents: 61664
diff changeset
   915
  from \<open>lower\<close>.
29706
10e6f2faa1e5 updated type class section
haftmann
parents: 29613
diff changeset
   916
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   917
  \<^descr> @{method intro_classes} repeatedly expands all class introduction rules of
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   918
  this theory. Note that this method usually needs not be named explicitly, as
61665
wenzelm
parents: 61664
diff changeset
   919
  it is already included in the default proof step (e.g.\ of \<^theory_text>\<open>proof\<close>). In
wenzelm
parents: 61664
diff changeset
   920
  particular, instantiation of trivial (syntactic) classes may be performed by
wenzelm
parents: 61664
diff changeset
   921
  a single ``\<^theory_text>\<open>..\<close>'' proof step.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   922
\<close>
26870
94bedbb34b92 misc reorganization;
wenzelm
parents: 26869
diff changeset
   923
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   924
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   925
subsection \<open>The class target\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   926
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   927
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   928
  %FIXME check
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   929
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   930
  A named context may refer to a locale (cf.\ \secref{sec:target}). If this
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   931
  locale is also a class \<open>c\<close>, apart from the common locale target behaviour
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   932
  the following happens.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   933
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   934
    \<^item> Local constant declarations \<open>g[\<alpha>]\<close> referring to the local type parameter
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   935
    \<open>\<alpha>\<close> and local parameters \<open>f[\<alpha>]\<close> are accompanied by theory-level constants
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   936
    \<open>g[?\<alpha> :: c]\<close> referring to theory-level class operations \<open>f[?\<alpha> :: c]\<close>.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   937
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   938
    \<^item> Local theorem bindings are lifted as are assumptions.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   939
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   940
    \<^item> Local syntax refers to local operations \<open>g[\<alpha>]\<close> and global operations
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   941
    \<open>g[?\<alpha> :: c]\<close> uniformly. Type inference resolves ambiguities. In rare
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   942
    cases, manual type annotations are needed.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   943
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   944
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   945
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   946
subsection \<open>Co-regularity of type classes and arities\<close>
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   947
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   948
text \<open>
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   949
  The class relation together with the collection of type-constructor arities
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   950
  must obey the principle of \<^emph>\<open>co-regularity\<close> as defined below.
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   951
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   952
  \<^medskip>
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   953
  For the subsequent formulation of co-regularity we assume that the class
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   954
  relation is closed by transitivity and reflexivity. Moreover the collection
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   955
  of arities \<open>t :: (\<^vec>s)c\<close> is completed such that \<open>t :: (\<^vec>s)c\<close> and
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   956
  \<open>c \<subseteq> c'\<close> implies \<open>t :: (\<^vec>s)c'\<close> for all such declarations.
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   957
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   958
  Treating sorts as finite sets of classes (meaning the intersection), the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   959
  class relation \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close> is extended to sorts as follows:
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   960
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   961
    \<open>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\<close>
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   962
  \]
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   963
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   964
  This relation on sorts is further extended to tuples of sorts (of the same
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   965
  length) in the component-wise way.
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   966
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   967
  \<^medskip>
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   968
  Co-regularity of the class relation together with the arities relation
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   969
  means:
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   970
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   971
    \<open>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\<close>
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   972
  \]
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   973
  for all such arities. In other words, whenever the result classes of some
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   974
  type-constructor arities are related, then the argument sorts need to be
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   975
  related in the same way.
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   976
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   977
  \<^medskip>
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   978
  Co-regularity is a very fundamental property of the order-sorted algebra of
71567
9a29e883a934 tuned documentation, based on hints by Pedro Sánchez Terraf;
wenzelm
parents: 71166
diff changeset
   979
  types. For example, it entails principal types and most general unifiers,
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
   980
  e.g.\ see @{cite "nipkow-prehofer"}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   981
\<close>
37768
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   982
e86221f3bac7 moved co-regularity to class section; avoid duplicated class_deps
haftmann
parents: 37112
diff changeset
   983
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   984
section \<open>Overloaded constant definitions \label{sec:overloading}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
   985
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   986
text \<open>
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   987
  Definitions essentially express abbreviations within the logic. The simplest
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   988
  form of a definition is \<open>c :: \<sigma> \<equiv> t\<close>, where \<open>c\<close> is a new constant and \<open>t\<close> is
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   989
  a closed term that does not mention \<open>c\<close>. Moreover, so-called \<^emph>\<open>hidden
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   990
  polymorphism\<close> is excluded: all type variables in \<open>t\<close> need to occur in its
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   991
  type \<open>\<sigma>\<close>.
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   992
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   993
  \<^emph>\<open>Overloading\<close> means that a constant being declared as \<open>c :: \<alpha> decl\<close> may be
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   994
  defined separately on type instances \<open>c :: (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n)\<kappa> decl\<close> for each
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   995
  type constructor \<open>\<kappa>\<close>. At most occasions overloading will be used in a
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   996
  Haskell-like fashion together with type classes by means of \<^theory_text>\<open>instantiation\<close>
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   997
  (see \secref{sec:class}). Sometimes low-level overloading is desirable; this
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   998
  is supported by \<^theory_text>\<open>consts\<close> and \<^theory_text>\<open>overloading\<close> explained below.
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
   999
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1000
  The right-hand side of overloaded definitions may mention overloaded
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1001
  constants recursively at type instances corresponding to the immediate
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1002
  argument types \<open>\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n\<close>. Incomplete specification patterns impose
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1003
  global constraints on all occurrences. E.g.\ \<open>d :: \<alpha> \<times> \<alpha>\<close> on the left-hand
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1004
  side means that all corresponding occurrences on some right-hand side need
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1005
  to be an instance of this, and general \<open>d :: \<alpha> \<times> \<beta>\<close> will be disallowed. Full
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1006
  details are given by Kun\v{c}ar @{cite "Kuncar:2015"}.
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1007
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1008
  \<^medskip>
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1009
  The \<^theory_text>\<open>consts\<close> command and the \<^theory_text>\<open>overloading\<close> target provide a convenient
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1010
  interface for end-users. Regular specification elements such as @{command
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1011
  definition}, @{command inductive}, @{command function} may be used in the
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1012
  body. It is also possible to use \<^theory_text>\<open>consts c :: \<sigma>\<close> with later \<^theory_text>\<open>overloading c
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1013
  \<equiv> c :: \<sigma>\<close> to keep the declaration and definition of a constant separate.
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1014
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
  1015
  \begin{matharray}{rcl}
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1016
    @{command_def "consts"} & : & \<open>theory \<rightarrow> theory\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1017
    @{command_def "overloading"} & : & \<open>theory \<rightarrow> local_theory\<close> \\
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
  1018
  \end{matharray}
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
  1019
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1020
  \<^rail>\<open>
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1021
    @@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +)
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1022
    ;
42704
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42651
diff changeset
  1023
    @@{command overloading} ( spec + ) @'begin'
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42651
diff changeset
  1024
    ;
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1025
    spec: @{syntax name} ( '\<equiv>' | '==' ) @{syntax term} ( '(' @'unchecked' ')' )?
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1026
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1027
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1028
  \<^descr> \<^theory_text>\<open>consts c :: \<sigma>\<close> declares constant \<open>c\<close> to have any instance of type scheme
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1029
  \<open>\<sigma>\<close>. The optional mixfix annotations may attach concrete syntax to the
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1030
  constants declared.
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1031
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1032
  \<^descr> \<^theory_text>\<open>overloading x\<^sub>1 \<equiv> c\<^sub>1 :: \<tau>\<^sub>1 \<dots> x\<^sub>n \<equiv> c\<^sub>n :: \<tau>\<^sub>n begin \<dots> end\<close> defines
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1033
  a theory target (cf.\ \secref{sec:target}) which allows to specify already
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1034
  declared constants via definitions in the body. These are identified by an
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1035
  explicitly given mapping from variable names \<open>x\<^sub>i\<close> to constants \<open>c\<^sub>i\<close> at
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1036
  particular type instances. The definitions themselves are established using
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1037
  common specification tools, using the names \<open>x\<^sub>i\<close> as reference to the
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1038
  corresponding constants.
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1039
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1040
  Option \<^theory_text>\<open>(unchecked)\<close> disables global dependency checks for the
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1041
  corresponding definition, which is occasionally useful for exotic
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1042
  overloading; this is a form of axiomatic specification. It is at the
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1043
  discretion of the user to avoid malformed theory specifications!
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1044
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1045
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1046
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1047
subsubsection \<open>Example\<close>
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1048
62173
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1049
consts Length :: "'a \<Rightarrow> nat"
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1050
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1051
overloading
62173
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1052
  Length\<^sub>0 \<equiv> "Length :: unit \<Rightarrow> nat"
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1053
  Length\<^sub>1 \<equiv> "Length :: 'a \<times> unit \<Rightarrow> nat"
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1054
  Length\<^sub>2 \<equiv> "Length :: 'a \<times> 'b \<times> unit \<Rightarrow> nat"
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1055
  Length\<^sub>3 \<equiv> "Length :: 'a \<times> 'b \<times> 'c \<times> unit \<Rightarrow> nat"
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1056
begin
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1057
62173
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1058
fun Length\<^sub>0 :: "unit \<Rightarrow> nat" where "Length\<^sub>0 () = 0"
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1059
fun Length\<^sub>1 :: "'a \<times> unit \<Rightarrow> nat" where "Length\<^sub>1 (a, ()) = 1"
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1060
fun Length\<^sub>2 :: "'a \<times> 'b \<times> unit \<Rightarrow> nat" where "Length\<^sub>2 (a, b, ()) = 2"
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1061
fun Length\<^sub>3 :: "'a \<times> 'b \<times> 'c \<times> unit \<Rightarrow> nat" where "Length\<^sub>3 (a, b, c, ()) = 3"
62172
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1062
7eaeae127955 updated section on "Overloaded constant definitions";
wenzelm
parents: 61891
diff changeset
  1063
end
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1064
62173
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1065
lemma "Length (a, b, c, ()) = 3" by simp
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1066
lemma "Length ((a, b), (c, d), ()) = 2" by simp
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1067
lemma "Length ((a, b, c, d, e), ()) = 1" by simp
a817e4335a31 clarified example;
wenzelm
parents: 62172
diff changeset
  1068
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1069
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1070
section \<open>Incorporating ML code \label{sec:ML}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1071
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1072
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1073
  \begin{matharray}{rcl}
62862
007c454d0d0f more uniform ML file commands;
wenzelm
parents: 62173
diff changeset
  1074
    @{command_def "SML_file"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1075
    @{command_def "SML_file_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1076
    @{command_def "SML_file_no_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1077
    @{command_def "ML_file"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1078
    @{command_def "ML_file_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1079
    @{command_def "ML_file_no_debug"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1080
    @{command_def "ML"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
68276
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1081
    @{command_def "ML_export"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1082
    @{command_def "ML_prf"} & : & \<open>proof \<rightarrow> proof\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1083
    @{command_def "ML_val"} & : & \<open>any \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1084
    @{command_def "ML_command"} & : & \<open>any \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1085
    @{command_def "setup"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1086
    @{command_def "local_setup"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1087
    @{command_def "attribute_setup"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28758
diff changeset
  1088
  \end{matharray}
57478
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1089
  \begin{tabular}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1090
    @{attribute_def ML_print_depth} & : & \<open>attribute\<close> & default 10 \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1091
    @{attribute_def ML_source_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1092
    @{attribute_def ML_debugger} & : & \<open>attribute\<close> & default \<open>false\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1093
    @{attribute_def ML_exception_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1094
    @{attribute_def ML_exception_debugger} & : & \<open>attribute\<close> & default \<open>false\<close> \\
70260
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1095
    @{attribute_def ML_environment} & : & \<open>attribute\<close> & default \<open>Isabelle\<close> \\
57478
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1096
  \end{tabular}
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28758
diff changeset
  1097
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1098
  \<^rail>\<open>
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1099
    (@@{command SML_file} |
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1100
      @@{command SML_file_debug} |
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1101
      @@{command SML_file_no_debug} |
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1102
      @@{command ML_file} |
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1103
      @@{command ML_file_debug} |
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
  1104
      @@{command ML_file_no_debug}) @{syntax name} ';'?
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1105
    ;
68276
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1106
    (@@{command ML} | @@{command ML_export} | @@{command ML_prf} |
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1107
      @@{command ML_val} | @@{command ML_command} | @@{command setup} |
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1108
      @@{command local_setup}) @{syntax text}
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1109
    ;
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 59781
diff changeset
  1110
    @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}?
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1111
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1112
61665
wenzelm
parents: 61664
diff changeset
  1113
  \<^descr> \<^theory_text>\<open>SML_file name\<close> reads and evaluates the given Standard ML file. Top-level
62862
007c454d0d0f more uniform ML file commands;
wenzelm
parents: 62173
diff changeset
  1114
  SML bindings are stored within the (global or local) theory context; the
007c454d0d0f more uniform ML file commands;
wenzelm
parents: 62173
diff changeset
  1115
  initial environment is restricted to the Standard ML implementation of
007c454d0d0f more uniform ML file commands;
wenzelm
parents: 62173
diff changeset
  1116
  Poly/ML, without the many add-ons of Isabelle/ML. Multiple \<^theory_text>\<open>SML_file\<close>
007c454d0d0f more uniform ML file commands;
wenzelm
parents: 62173
diff changeset
  1117
  commands may be used to build larger Standard ML projects, independently of
007c454d0d0f more uniform ML file commands;
wenzelm
parents: 62173
diff changeset
  1118
  the regular Isabelle/ML environment.
56275
600f432ab556 added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents: 55837
diff changeset
  1119
61665
wenzelm
parents: 61664
diff changeset
  1120
  \<^descr> \<^theory_text>\<open>ML_file name\<close> reads and evaluates the given ML file. The current theory
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1121
  context is passed down to the ML toplevel and may be modified, using \<^ML>\<open>Context.>>\<close> or derived ML commands. Top-level ML bindings are stored
61665
wenzelm
parents: 61664
diff changeset
  1122
  within the (global or local) theory context.
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1123
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1124
  \<^descr> \<^theory_text>\<open>SML_file_debug\<close>, \<^theory_text>\<open>SML_file_no_debug\<close>, \<^theory_text>\<open>ML_file_debug\<close>, and
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1125
  \<^theory_text>\<open>ML_file_no_debug\<close> change the @{attribute ML_debugger} option locally while
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1126
  the given file is compiled.
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1127
68276
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1128
  \<^descr> \<^theory_text>\<open>ML\<close> is similar to \<^theory_text>\<open>ML_file\<close>, but evaluates directly the given \<open>text\<close>.
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1129
  Top-level ML bindings are stored within the (global or local) theory
61665
wenzelm
parents: 61664
diff changeset
  1130
  context.
28281
132456af0731 added ML_prf;
wenzelm
parents: 28114
diff changeset
  1131
68276
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1132
  \<^descr> \<^theory_text>\<open>ML_export\<close> is similar to \<^theory_text>\<open>ML\<close>, but the resulting toplevel bindings are
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1133
  exported to the global bootstrap environment of the ML process --- it has
68278
wenzelm
parents: 68276
diff changeset
  1134
  a lasting effect that cannot be retracted. This allows ML evaluation
68276
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1135
  without a formal theory context, e.g. for command-line tools via @{tool
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1136
  process} @{cite "isabelle-system"}.
cbee43ff4ceb added command 'ML_export';
wenzelm
parents: 67764
diff changeset
  1137
61665
wenzelm
parents: 61664
diff changeset
  1138
  \<^descr> \<^theory_text>\<open>ML_prf\<close> is analogous to \<^theory_text>\<open>ML\<close> but works within a proof context.
wenzelm
parents: 61664
diff changeset
  1139
  Top-level ML bindings are stored within the proof context in a purely
wenzelm
parents: 61664
diff changeset
  1140
  sequential fashion, disregarding the nested proof structure. ML bindings
wenzelm
parents: 61664
diff changeset
  1141
  introduced by \<^theory_text>\<open>ML_prf\<close> are discarded at the end of the proof.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1142
61665
wenzelm
parents: 61664
diff changeset
  1143
  \<^descr> \<^theory_text>\<open>ML_val\<close> and \<^theory_text>\<open>ML_command\<close> are diagnostic versions of \<^theory_text>\<open>ML\<close>, which means
wenzelm
parents: 61664
diff changeset
  1144
  that the context may not be updated. \<^theory_text>\<open>ML_val\<close> echos the bindings produced
wenzelm
parents: 61664
diff changeset
  1145
  at the ML toplevel, but \<^theory_text>\<open>ML_command\<close> is silent.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
  1146
61665
wenzelm
parents: 61664
diff changeset
  1147
  \<^descr> \<^theory_text>\<open>setup "text"\<close> changes the current theory context by applying \<open>text\<close>,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1148
  which refers to an ML expression of type \<^ML_type>\<open>theory -> theory\<close>. This
61665
wenzelm
parents: 61664
diff changeset
  1149
  enables to initialize any object-logic specific tools and packages written
wenzelm
parents: 61664
diff changeset
  1150
  in ML, for example.
30461
00323c45ea83 added 'local_setup' command;
wenzelm
parents: 30242
diff changeset
  1151
61665
wenzelm
parents: 61664
diff changeset
  1152
  \<^descr> \<^theory_text>\<open>local_setup\<close> is similar to \<^theory_text>\<open>setup\<close> for a local theory context, and an
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1153
  ML expression of type \<^ML_type>\<open>local_theory -> local_theory\<close>. This allows
61665
wenzelm
parents: 61664
diff changeset
  1154
  to invoke local theory specification packages without going through concrete
wenzelm
parents: 61664
diff changeset
  1155
  outer syntax, for example.
28758
4ce896a30f88 added bind_thm, bind_thms;
wenzelm
parents: 28757
diff changeset
  1156
61665
wenzelm
parents: 61664
diff changeset
  1157
  \<^descr> \<^theory_text>\<open>attribute_setup name = "text" description\<close> defines an attribute in the
wenzelm
parents: 61664
diff changeset
  1158
  current context. The given \<open>text\<close> has to be an ML expression of type
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1159
  \<^ML_type>\<open>attribute context_parser\<close>, cf.\ basic parsers defined in
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1160
  structure \<^ML_structure>\<open>Args\<close> and \<^ML_structure>\<open>Attrib\<close>.
30526
7f9a9ec1c94d added 'attribute_setup' command;
wenzelm
parents: 30461
diff changeset
  1161
7f9a9ec1c94d added 'attribute_setup' command;
wenzelm
parents: 30461
diff changeset
  1162
  In principle, attributes can operate both on a given theorem and the
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1163
  implicit context, although in practice only one is modified and the other
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1164
  serves as parameter. Here are examples for these two cases:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1165
\<close>
30526
7f9a9ec1c94d added 'attribute_setup' command;
wenzelm
parents: 30461
diff changeset
  1166
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59901
diff changeset
  1167
(*<*)experiment begin(*>*)
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1168
        attribute_setup my_rule =
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1169
          \<open>Attrib.thms >> (fn ths =>
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1170
            Thm.rule_attribute ths
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1171
              (fn context: Context.generic => fn th: thm =>
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1172
                let val th' = th OF ths
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1173
                in th' end))\<close>
30526
7f9a9ec1c94d added 'attribute_setup' command;
wenzelm
parents: 30461
diff changeset
  1174
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1175
        attribute_setup my_declaration =
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1176
          \<open>Attrib.thms >> (fn ths =>
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1177
            Thm.declaration_attribute
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1178
              (fn th: thm => fn context: Context.generic =>
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1179
                let val context' = context
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1180
                in context' end))\<close>
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59901
diff changeset
  1181
(*<*)end(*>*)
30526
7f9a9ec1c94d added 'attribute_setup' command;
wenzelm
parents: 30461
diff changeset
  1182
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1183
text \<open>
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1184
  \<^descr> @{attribute ML_print_depth} controls the printing depth of the ML toplevel
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1185
  pretty printer. Typically the limit should be less than 10. Bigger values
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1186
  such as 100--1000 are occasionally useful for debugging.
57478
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1187
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1188
  \<^descr> @{attribute ML_source_trace} indicates whether the source text that is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1189
  given to the ML compiler should be output: it shows the raw Standard ML
57478
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1190
  after expansion of Isabelle/ML antiquotations.
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1191
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1192
  \<^descr> @{attribute ML_debugger} controls compilation of sources with or without
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1193
  debugging information. The global system option @{system_option_ref
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1194
  ML_debugger} does the same when building a session image. It is also
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1195
  possible use commands like \<^theory_text>\<open>ML_file_debug\<close> etc. The ML debugger is
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1196
  explained further in @{cite "isabelle-jedit"}.
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1197
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1198
  \<^descr> @{attribute ML_exception_trace} indicates whether the ML run-time system
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1199
  should print a detailed stack trace on exceptions. The result is dependent
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1200
  on various ML compiler optimizations. The boundary for the exception trace
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1201
  is the current Isar command transactions: it is occasionally better to
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1202
  insert the combinator \<^ML>\<open>Runtime.exn_trace\<close> into ML code for debugging
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1203
  @{cite "isabelle-implementation"}, closer to the point where it actually
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1204
  happens.
57478
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1205
62903
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1206
  \<^descr> @{attribute ML_exception_debugger} controls detailed exception trace via
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1207
  the Poly/ML debugger, at the cost of extra compile-time and run-time
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1208
  overhead. Relevant ML modules need to be compiled beforehand with debugging
adcce7b8d8ba updated documentation;
wenzelm
parents: 62862
diff changeset
  1209
  enabled, see @{attribute ML_debugger} above.
70260
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1210
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1211
  \<^descr> @{attribute ML_environment} determines the named ML environment for
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1212
  toplevel declarations, e.g.\ in command \<^theory_text>\<open>ML\<close> or \<^theory_text>\<open>ML_file\<close>. The following
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1213
  ML environments are predefined in Isabelle/Pure:
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1214
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1215
    \<^item> \<open>Isabelle\<close> for Isabelle/ML. It contains all modules of Isabelle/Pure and
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1216
    further add-ons, e.g. material from Isabelle/HOL.
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1217
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1218
    \<^item> \<open>SML\<close> for official Standard ML. It contains only the initial basis
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1219
    according to \<^url>\<open>http://sml-family.org/Basis/overview.html\<close>.
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1220
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1221
  The Isabelle/ML function \<^ML>\<open>ML_Env.setup\<close> defines a new ML environment.
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1222
  This is useful to incorporate big SML projects in an isolated name space,
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1223
  possibly with variations on ML syntax; the existing setup of
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1224
  \<^ML>\<open>ML_Env.SML_operations\<close> follows the official standard.
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1225
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1226
  It is also possible to move toplevel bindings between ML environments, using
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1227
  a notation with ``\<open>>\<close>'' as separator. For example:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1228
\<close>
57478
fa14d60a8cca more on ML options;
wenzelm
parents: 56594
diff changeset
  1229
70260
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1230
(*<*)experiment begin(*>*)
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1231
        declare [[ML_environment = "Isabelle>SML"]]
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1232
        ML \<open>val println = writeln\<close>
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1233
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1234
        declare [[ML_environment = "SML"]]
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1235
        ML \<open>println "test"\<close>
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1236
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1237
        declare [[ML_environment = "Isabelle"]]
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1238
        ML \<open>ML \<open>println\<close> (*bad*) handle ERROR msg => warning msg\<close>
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1239
(*<*)end(*>*)
22cfcfcadd8b more documentation;
wenzelm
parents: 70058
diff changeset
  1240
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1241
70261
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1242
section \<open>Generated files and exported files\<close>
66757
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1243
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1244
text \<open>
70057
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1245
  Write access to the physical file-system is incompatible with the stateless
70261
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1246
  model of processing Isabelle documents. To avoid bad effects, the following
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1247
  concepts for abstract file-management are provided by Isabelle:
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1248
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1249
    \<^descr>[Generated files] are stored within the theory context in Isabelle/ML.
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1250
    This allows to operate on the content in Isabelle/ML, e.g. via the command
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1251
    @{command compile_generated_files}.
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1252
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1253
    \<^descr>[Exported files] are stored within the session database in
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1254
    Isabelle/Scala. This allows to deliver artefacts to external tools, see
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1255
    also @{cite "isabelle-system"} for session \<^verbatim>\<open>ROOT\<close> declaration
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1256
    \<^theory_text>\<open>export_files\<close>, and @{tool build} option \<^verbatim>\<open>-e\<close>.
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1257
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1258
  A notable example is the command @{command_ref export_code}
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1259
  (\chref{ch:export-code}): it uses both concepts simultaneously.
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1260
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1261
  File names are hierarchically structured, using a slash as separator. The
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1262
  (long) theory name is used as a prefix: the resulting name needs to be
efbdfcaa6258 clarified documentation;
wenzelm
parents: 70260
diff changeset
  1263
  globally unique.
70057
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1264
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1265
  \begin{matharray}{rcll}
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1266
    @{command_def "generate_file"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1267
    @{command_def "export_generated_files"} & : & \<open>context \<rightarrow>\<close> \\
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1268
    @{command_def "compile_generated_files"} & : & \<open>context \<rightarrow>\<close> \\
66757
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1269
    @{command_def "external_file"} & : & \<open>any \<rightarrow> any\<close> \\
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1270
  \end{matharray}
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1271
70057
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1272
  \<^rail>\<open>
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1273
    @@{command generate_file} path '=' content
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1274
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1275
    path: @{syntax embedded}
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1276
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1277
    content: @{syntax embedded}
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1278
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1279
    @@{command export_generated_files} (files_in_theory + @'and')
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1280
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1281
    files_in_theory: (@'_' | (path+)) (('(' @'in' @{syntax name} ')')?)
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1282
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1283
    @@{command compile_generated_files} (files_in_theory + @'and') \<newline>
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1284
      (@'external_files' (external_files + @'and'))? \<newline>
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1285
      (@'export_files' (export_files + @'and'))? \<newline>
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1286
      (@'export_prefix' path)?
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1287
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1288
    external_files: (path+) (('(' @'in' path ')')?)
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1289
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1290
    export_files: (path+) (executable?)
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1291
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1292
    executable: '(' ('exe' | 'executable') ')'
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1293
    ;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1294
    @@{command external_file} @{syntax name} ';'?
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1295
  \<close>
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1296
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1297
  \<^descr> \<^theory_text>\<open>generate_file path = content\<close> augments the table of generated files
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1298
  within the current theory by a new entry: duplicates are not allowed. The
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1299
  name extension determines a pre-existent file-type; the \<open>content\<close> is a
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1300
  string that is preprocessed according to rules of this file-type.
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1301
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1302
  For example, Isabelle/Pure supports \<^verbatim>\<open>.hs\<close> as file-type for Haskell:
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1303
  embedded cartouches are evaluated as Isabelle/ML expressions of type
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1304
  \<^ML_type>\<open>string\<close>, the result is inlined in Haskell string syntax.
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1305
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1306
  \<^descr> \<^theory_text>\<open>export_generated_files paths (in thy)\<close> retrieves named generated files
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1307
  from the given theory (that needs be reachable via imports of the current
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1308
  one). By default, the current theory node is used. Using ``\<^verbatim>\<open>_\<close>''
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1309
  (underscore) instead of explicit path names refers to \emph{all} files of a
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1310
  theory node.
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1311
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1312
  The overall list of files is prefixed with the respective (long) theory name
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1313
  and exported to the session database. In Isabelle/jEdit the result can be
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1314
  browsed via the virtual file-system with prefix ``\<^verbatim>\<open>isabelle-export:\<close>''
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1315
  (using the regular file-browser).
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1316
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1317
  \<^descr> \<^theory_text>\<open>compile_generated_files paths (in thy) where compile_body\<close> retrieves
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1318
  named generated files as for \<^theory_text>\<open>export_generated_files\<close> and writes them into
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1319
  a temporary directory, such that the \<open>compile_body\<close> may operate on them as
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1320
  an ML function of type \<^ML_type>\<open>Path.T -> unit\<close>. This may create further
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1321
  files, e.g.\ executables produced by a compiler that is invoked as external
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1322
  process (e.g.\ via \<^ML>\<open>Isabelle_System.bash\<close>), or any other files.
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1323
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1324
  The option ``\<^theory_text>\<open>external_files paths (in base_dir)\<close>'' copies files from the
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1325
  physical file-system into the temporary directory, \emph{before} invoking
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1326
  \<open>compile_body\<close>. The \<open>base_dir\<close> prefix is removed from each of the \<open>paths\<close>,
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1327
  but the remaining sub-directory structure is reconstructed in the target
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1328
  directory.
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1329
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1330
  The option ``\<^theory_text>\<open>export_files paths\<close>'' exports the specified files from the
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1331
  temporary directory to the session database, \emph{after} invoking
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1332
  \<open>compile_body\<close>. Entries may be decorated with ``\<^theory_text>\<open>(exe)\<close>'' to say that it is
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1333
  a platform-specific executable program: the executable file-attribute will
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1334
  be set, and on Windows the \<^verbatim>\<open>.exe\<close> file-extension will be included;
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1335
  ``\<^theory_text>\<open>(executable)\<close>'' only refers to the file-attribute, without special
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1336
  treatment of the \<^verbatim>\<open>.exe\<close> extension.
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1337
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1338
  The option ``\<^theory_text>\<open>export_prefix path\<close>'' specifies an extra path prefix for all
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1339
  exports of \<^theory_text>\<open>export_files\<close> above.
66757
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1340
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1341
  \<^descr> \<^theory_text>\<open>external_file name\<close> declares the formal dependency on the given file
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1342
  name, such that the Isabelle build process knows about it (see also @{cite
70057
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1343
  "isabelle-system"}). This is required for any files mentioned in
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1344
  \<^theory_text>\<open>compile_generated_files / external_files\<close> above, in order to document
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1345
  source dependencies properly. It is also possible to use \<^theory_text>\<open>external_file\<close>
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1346
  alone, e.g.\ when other Isabelle/ML tools use \<^ML>\<open>File.read\<close>, without
0403b5127da1 documentation for generated files;
wenzelm
parents: 69962
diff changeset
  1347
  specific management of content by the Prover IDE.
66757
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1348
\<close>
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1349
e32750d7acb4 added command 'external_file';
wenzelm
parents: 66248
diff changeset
  1350
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1351
section \<open>Primitive specification elements\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1352
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1353
subsection \<open>Sorts\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1354
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1355
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1356
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1357
    @{command_def "default_sort"} & : & \<open>local_theory \<rightarrow> local_theory\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1358
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1359
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1360
  \<^rail>\<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
  1361
    @@{command default_sort} @{syntax sort}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1362
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1363
61665
wenzelm
parents: 61664
diff changeset
  1364
  \<^descr> \<^theory_text>\<open>default_sort s\<close> makes sort \<open>s\<close> the new default sort for any type
wenzelm
parents: 61664
diff changeset
  1365
  variable that is given explicitly in the text, but lacks a sort constraint
wenzelm
parents: 61664
diff changeset
  1366
  (wrt.\ the current context). Type variables generated by type inference are
wenzelm
parents: 61664
diff changeset
  1367
  not affected.
28767
f09ceb800d00 minor tuning (according to old ref manual);
wenzelm
parents: 28762
diff changeset
  1368
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1369
  Usually the default sort is only changed when defining a new object-logic.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1370
  For example, the default sort in Isabelle/HOL is \<^class>\<open>type\<close>, the class of
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1371
  all HOL types.
28767
f09ceb800d00 minor tuning (according to old ref manual);
wenzelm
parents: 28762
diff changeset
  1372
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1373
  When merging theories, the default sorts of the parents are logically
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1374
  intersected, i.e.\ the representations as lists of classes are joined.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1375
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1376
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1377
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1378
subsection \<open>Types \label{sec:types-pure}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1379
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1380
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1381
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1382
    @{command_def "type_synonym"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1383
    @{command_def "typedecl"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1384
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1385
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1386
  \<^rail>\<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
  1387
    @@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?)
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1388
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
  1389
    @@{command typedecl} @{syntax typespec} @{syntax mixfix}?
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1390
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1391
61665
wenzelm
parents: 61664
diff changeset
  1392
  \<^descr> \<^theory_text>\<open>type_synonym (\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = \<tau>\<close> introduces a \<^emph>\<open>type synonym\<close> \<open>(\<alpha>\<^sub>1, \<dots>,
wenzelm
parents: 61664
diff changeset
  1393
  \<alpha>\<^sub>n) t\<close> for the existing type \<open>\<tau>\<close>. Unlike the semantic type definitions in
wenzelm
parents: 61664
diff changeset
  1394
  Isabelle/HOL, type synonyms are merely syntactic abbreviations without any
wenzelm
parents: 61664
diff changeset
  1395
  logical significance. Internally, type synonyms are fully expanded.
65491
7fb81fa1d668 more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents: 63680
diff changeset
  1396
61665
wenzelm
parents: 61664
diff changeset
  1397
  \<^descr> \<^theory_text>\<open>typedecl (\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t\<close> declares a new type constructor \<open>t\<close>. If the
wenzelm
parents: 61664
diff changeset
  1398
  object-logic defines a base sort \<open>s\<close>, then the constructor is declared to
wenzelm
parents: 61664
diff changeset
  1399
  operate on that, via the axiomatic type-class instance \<open>t :: (s, \<dots>, s)s\<close>.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1400
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
  1401
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
  1402
  \begin{warn}
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1403
    If you introduce a new type axiomatically, i.e.\ via @{command_ref
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1404
    typedecl} and @{command_ref axiomatization}
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1405
    (\secref{sec:axiomatizations}), the minimum requirement is that it has a
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1406
    non-empty model, to avoid immediate collapse of the logical environment.
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1407
    Moreover, one needs to demonstrate that the interpretation of such
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1408
    free-form axiomatizations can coexist with other axiomatization schemes
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1409
    for types, notably @{command_def typedef} in Isabelle/HOL
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1410
    (\secref{sec:hol-typedef}), or any other extension that people might have
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1411
    introduced elsewhere.
57487
7806a74c54ac misc tuning and clarification;
wenzelm
parents: 57480
diff changeset
  1412
  \end{warn}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1413
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1414
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1415
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1416
section \<open>Naming existing theorems \label{sec:theorems}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1417
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1418
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1419
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1420
    @{command_def "lemmas"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1421
    @{command_def "named_theorems"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1422
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1423
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1424
  \<^rail>\<open>
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62965
diff changeset
  1425
    @@{command lemmas} (@{syntax thmdef}? @{syntax thms} + @'and')
59785
4e6ab5831cc0 clarified syntax category "fixes";
wenzelm
parents: 59783
diff changeset
  1426
      @{syntax for_fixes}
57946
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57941
diff changeset
  1427
    ;
59785
4e6ab5831cc0 clarified syntax category "fixes";
wenzelm
parents: 59783
diff changeset
  1428
    @@{command named_theorems} (@{syntax name} @{syntax text}? + @'and')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1429
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1430
61665
wenzelm
parents: 61664
diff changeset
  1431
  \<^descr> \<^theory_text>\<open>lemmas a = b\<^sub>1 \<dots> b\<^sub>n\<close>~@{keyword_def "for"}~\<open>x\<^sub>1 \<dots> x\<^sub>m\<close> evaluates given
wenzelm
parents: 61664
diff changeset
  1432
  facts (with attributes) in the current context, which may be augmented by
wenzelm
parents: 61664
diff changeset
  1433
  local variables. Results are standardized before being stored, i.e.\
wenzelm
parents: 61664
diff changeset
  1434
  schematic variables are renamed to enforce index \<open>0\<close> uniformly.
45600
1bbbac9a0cb0 'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents: 42936
diff changeset
  1435
61665
wenzelm
parents: 61664
diff changeset
  1436
  \<^descr> \<^theory_text>\<open>named_theorems name description\<close> declares a dynamic fact within the
wenzelm
parents: 61664
diff changeset
  1437
  context. The same \<open>name\<close> is used to define an attribute with the usual
wenzelm
parents: 61664
diff changeset
  1438
  \<open>add\<close>/\<open>del\<close> syntax (e.g.\ see \secref{sec:simp-rules}) to maintain the
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1439
  content incrementally, in canonical declaration order of the text structure.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1440
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1441
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1442
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1443
section \<open>Oracles \label{sec:oracles}\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1444
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1445
text \<open>
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
  1446
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1447
    @{command_def "oracle"} & : & \<open>theory \<rightarrow> theory\<close> & (axiomatic!) \\
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1448
    @{command_def "thm_oracles"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
47483
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
  1449
  \end{matharray}
3f704717a67f more uniform outline;
wenzelm
parents: 47482
diff changeset
  1450
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1451
  Oracles allow Isabelle to take advantage of external reasoners such as
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1452
  arithmetic decision procedures, model checkers, fast tautology checkers or
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1453
  computer algebra systems. Invoked as an oracle, an external reasoner can
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1454
  create arbitrary Isabelle theorems.
28756
529798e71924 tuned section "Oracles";
wenzelm
parents: 28745
diff changeset
  1455
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1456
  It is the responsibility of the user to ensure that the external reasoner is
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1457
  as trustworthy as the application requires. Another typical source of errors
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1458
  is the linkup between Isabelle and the external tool, not just its concrete
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1459
  implementation, but also the required translation between two different
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1460
  logical environments.
28756
529798e71924 tuned section "Oracles";
wenzelm
parents: 28745
diff changeset
  1461
529798e71924 tuned section "Oracles";
wenzelm
parents: 28745
diff changeset
  1462
  Isabelle merely guarantees well-formedness of the propositions being
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1463
  asserted, and records within the internal derivation object how presumed
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1464
  theorems depend on unproven suppositions. This also includes implicit
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1465
  type-class reasoning via the order-sorted algebra of class relations and
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1466
  type arities (see also @{command_ref instantiation} and @{command_ref
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1467
  instance}).
28756
529798e71924 tuned section "Oracles";
wenzelm
parents: 28745
diff changeset
  1468
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1469
  \<^rail>\<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 41435
diff changeset
  1470
    @@{command oracle} @{syntax name} '=' @{syntax text}
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1471
    ;
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1472
    @@{command thm_oracles} @{syntax thms}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1473
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1474
61665
wenzelm
parents: 61664
diff changeset
  1475
  \<^descr> \<^theory_text>\<open>oracle name = "text"\<close> turns the given ML expression \<open>text\<close> of type
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1476
  \<^ML_text>\<open>'a -> cterm\<close> into an ML function of type \<^ML_text>\<open>'a -> thm\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1477
  which is bound to the global identifier \<^ML_text>\<open>name\<close>. This acts like an
61665
wenzelm
parents: 61664
diff changeset
  1478
  infinitary specification of axioms! Invoking the oracle only works within
wenzelm
parents: 61664
diff changeset
  1479
  the scope of the resulting theory.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1480
71926
bee83c9d3306 clarified sessions;
wenzelm
parents: 71567
diff changeset
  1481
  See \<^file>\<open>~~/src/HOL/Examples/Iff_Oracle.thy\<close> for a worked example of defining
bee83c9d3306 clarified sessions;
wenzelm
parents: 71567
diff changeset
  1482
  a new primitive rule as oracle, and turning it into a proof method.
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1483
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1484
  \<^descr> \<^theory_text>\<open>thm_oracles thms\<close> displays all oracles used in the internal derivation
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1485
  of the given theorems; this covers the full graph of transitive
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70261
diff changeset
  1486
  dependencies.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1487
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1488
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1489
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1490
section \<open>Name spaces\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1491
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1492
text \<open>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1493
  \begin{matharray}{rcl}
66248
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1494
    @{command_def "alias"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1495
    @{command_def "type_alias"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1496
    @{command_def "hide_class"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1497
    @{command_def "hide_type"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1498
    @{command_def "hide_const"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1499
    @{command_def "hide_fact"} & : & \<open>theory \<rightarrow> theory\<close> \\
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1500
  \end{matharray}
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1501
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1502
  \<^rail>\<open>
66248
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1503
    (@{command alias} | @{command type_alias}) @{syntax name} '=' @{syntax name}
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1504
    ;
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1505
    (@{command hide_class} | @{command hide_type} |
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1506
      @{command hide_const} | @{command hide_fact}) ('(' @'open' ')')? (@{syntax name} + )
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69018
diff changeset
  1507
  \<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1508
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1509
  Isabelle organizes any kind of name declarations (of types, constants,
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1510
  theorems etc.) by separate hierarchically structured name spaces. Normally
61665
wenzelm
parents: 61664
diff changeset
  1511
  the user does not have to control the behaviour of name spaces by hand, yet
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1512
  the following commands provide some way to do so.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1513
66248
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1514
  \<^descr> \<^theory_text>\<open>alias\<close> and \<^theory_text>\<open>type_alias\<close> introduce aliases for constants and type
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1515
  constructors, respectively. This allows adhoc changes to name-space
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1516
  accesses.
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1517
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1518
  \<^descr> \<^theory_text>\<open>type_alias b = c\<close> introduces an alias for an existing type constructor.
df85956228c2 added command 'alias' and 'type_alias';
wenzelm
parents: 65491
diff changeset
  1519
61665
wenzelm
parents: 61664
diff changeset
  1520
  \<^descr> \<^theory_text>\<open>hide_class names\<close> fully removes class declarations from a given name
wenzelm
parents: 61664
diff changeset
  1521
  space; with the \<open>(open)\<close> option, only the unqualified base name is hidden.
36177
8e0770d2e499 separate commands 'hide_class', 'hide_type', 'hide_const', 'hide_fact';
wenzelm
parents: 35681
diff changeset
  1522
61664
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1523
  Note that hiding name space accesses has no impact on logical declarations
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1524
  --- they remain valid internally. Entities that are no longer accessible to
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1525
  the user are printed with the special qualifier ``\<open>??\<close>'' prefixed to the
6099d48193d0 tuned whitespace;
wenzelm
parents: 61656
diff changeset
  1526
  full internal name.
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1527
61665
wenzelm
parents: 61664
diff changeset
  1528
  \<^descr> \<^theory_text>\<open>hide_type\<close>, \<^theory_text>\<open>hide_const\<close>, and \<^theory_text>\<open>hide_fact\<close> are similar to
wenzelm
parents: 61664
diff changeset
  1529
  \<^theory_text>\<open>hide_class\<close>, but hide types, constants, and facts, respectively.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1530
\<close>
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26870
diff changeset
  1531
26869
3bc332135aa7 added chapters for "Specifications" and "Proofs";
wenzelm
parents:
diff changeset
  1532
end