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