author | wenzelm |
Tue, 03 Dec 2019 16:40:04 +0100 | |
changeset 71222 | 2bc39c80a95d |
parent 69597 | ff784d5a5bfb |
permissions | -rw-r--r-- |
27063 | 1 |
theory Examples1 |
2 |
imports Examples |
|
3 |
begin |
|
57607
5ff0cf3f5f6f
discontinued presumable workarounds for extra inter-theory space, which are obsolete since 0e5fa27d3293;
wenzelm
parents:
48985
diff
changeset
|
4 |
|
61419 | 5 |
section \<open>Use of Locales in Theories and Proofs |
6 |
\label{sec:interpretation}\<close> |
|
27063 | 7 |
|
61419 | 8 |
text \<open> |
32983 | 9 |
Locales can be interpreted in the contexts of theories and |
27063 | 10 |
structured proofs. These interpretations are dynamic, too. |
32981 | 11 |
Conclusions of locales will be propagated to the current theory or |
12 |
the current proof context.% |
|
13 |
\footnote{Strictly speaking, only interpretation in theories is |
|
14 |
dynamic since it is not possible to change locales or the locale |
|
15 |
hierarchy from within a proof.} |
|
16 |
The focus of this section is on |
|
17 |
interpretation in theories, but we will also encounter |
|
18 |
interpretations in proofs, in |
|
19 |
Section~\ref{sec:local-interpretation}. |
|
30580
cc5a55d7a5be
Updated chapters 1-5 to locale reimplementation.
ballarin
parents:
29293
diff
changeset
|
20 |
|
69597 | 21 |
As an example, consider the type of integers \<^typ>\<open>int\<close>. The |
22 |
relation \<^term>\<open>(\<le>)\<close> is a total order over \<^typ>\<open>int\<close>. We start |
|
23 |
with the interpretation that \<^term>\<open>(\<le>)\<close> is a partial order. The |
|
32983 | 24 |
facilities of the interpretation command are explored gradually in |
25 |
three versions. |
|
61419 | 26 |
\<close> |
27063 | 27 |
|
28 |
||
61419 | 29 |
subsection \<open>First Version: Replacement of Parameters Only |
30 |
\label{sec:po-first}\<close> |
|
27063 | 31 |
|
61419 | 32 |
text \<open> |
32981 | 33 |
The command \isakeyword{interpretation} is for the interpretation of |
34 |
locale in theories. In the following example, the parameter of locale |
|
69597 | 35 |
\<open>partial_order\<close> is replaced by \<^term>\<open>(\<le>) :: int \<Rightarrow> int \<Rightarrow> |
36 |
bool\<close> and the locale instance is interpreted in the current |
|
61419 | 37 |
theory.\<close> |
27063 | 38 |
|
67399 | 39 |
interpretation %visible int: partial_order "(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool" |
61419 | 40 |
txt \<open>\normalsize |
32981 | 41 |
The argument of the command is a simple \emph{locale expression} |
42 |
consisting of the name of the interpreted locale, which is |
|
69505 | 43 |
preceded by the qualifier \<open>int:\<close> and succeeded by a |
32981 | 44 |
white-space-separated list of terms, which provide a full |
45 |
instantiation of the locale parameters. The parameters are referred |
|
46 |
to by order of declaration, which is also the order in which |
|
32983 | 47 |
\isakeyword{print\_locale} outputs them. The locale has only a |
48 |
single parameter, hence the list of instantiation terms is a |
|
49 |
singleton. |
|
32981 | 50 |
|
51 |
The command creates the goal |
|
30580
cc5a55d7a5be
Updated chapters 1-5 to locale reimplementation.
ballarin
parents:
29293
diff
changeset
|
52 |
@{subgoals [display]} which can be shown easily: |
61419 | 53 |
\<close> |
27063 | 54 |
by unfold_locales auto |
55 |
||
61419 | 56 |
text \<open>The effect of the command is that instances of all |
32981 | 57 |
conclusions of the locale are available in the theory, where names |
69597 | 58 |
are prefixed by the qualifier. For example, transitivity for \<^typ>\<open>int\<close> is named @{thm [source] int.trans} and is the following |
32981 | 59 |
theorem: |
32982 | 60 |
@{thm [display, indent=2] int.trans} |
69505 | 61 |
It is not possible to reference this theorem simply as \<open>trans\<close>. This prevents unwanted hiding of existing theorems of the |
61419 | 62 |
theory by an interpretation.\<close> |
27063 | 63 |
|
64 |
||
61419 | 65 |
subsection \<open>Second Version: Replacement of Definitions\<close> |
27063 | 66 |
|
61419 | 67 |
text \<open>Not only does the above interpretation qualify theorem names. |
69505 | 68 |
The prefix \<open>int\<close> is applied to all names introduced in locale |
32981 | 69 |
conclusions including names introduced in definitions. The |
69505 | 70 |
qualified name \<open>int.less\<close> is short for |
71 |
the interpretation of the definition, which is \<open>partial_order.less (\<le>)\<close>. |
|
32981 | 72 |
Qualified name and expanded form may be used almost |
73 |
interchangeably.% |
|
69597 | 74 |
\footnote{Since \<^term>\<open>(\<le>)\<close> is polymorphic, for \<open>partial_order.less (\<le>)\<close> a |
69505 | 75 |
more general type will be inferred than for \<open>int.less\<close> which |
69597 | 76 |
is over type \<^typ>\<open>int\<close>.} |
61701
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61566
diff
changeset
|
77 |
The former is preferred on output, as for example in the theorem |
32982 | 78 |
@{thm [source] int.less_le_trans}: @{thm [display, indent=2] |
79 |
int.less_le_trans} |
|
32981 | 80 |
Both notations for the strict order are not satisfactory. The |
69597 | 81 |
constant \<^term>\<open>(<)\<close> is the strict order for \<^typ>\<open>int\<close>. |
32981 | 82 |
In order to allow for the desired replacement, interpretation |
83 |
accepts \emph{equations} in addition to the parameter instantiation. |
|
84 |
These follow the locale expression and are indicated with the |
|
61566
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents:
61419
diff
changeset
|
85 |
keyword \isakeyword{rewrites}. This is the revised interpretation: |
61419 | 86 |
\<close> |
27063 | 87 |
end |