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