author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
27063 | 1 |
theory Examples3 |
32983 | 2 |
imports Examples |
27063 | 3 |
begin |
57607
5ff0cf3f5f6f
discontinued presumable workarounds for extra inter-theory space, which are obsolete since 0e5fa27d3293;
wenzelm
parents:
54703
diff
changeset
|
4 |
|
61419 | 5 |
subsection \<open>Third Version: Local Interpretation |
6 |
\label{sec:local-interpretation}\<close> |
|
27063 | 7 |
|
69597 | 8 |
text \<open>In the above example, the fact that \<^term>\<open>(\<le>)\<close> is a partial |
32982 | 9 |
order for the integers was used in the second goal to |
69505 | 10 |
discharge the premise in the definition of \<open>(\<sqsubset>)\<close>. In |
32981 | 11 |
general, proofs of the equations not only may involve definitions |
32983 | 12 |
from the interpreted locale but arbitrarily complex arguments in the |
46855 | 13 |
context of the locale. Therefore it would be convenient to have the |
14 |
interpreted locale conclusions temporarily available in the proof. |
|
32981 | 15 |
This can be achieved by a locale interpretation in the proof body. |
16 |
The command for local interpretations is \isakeyword{interpret}. We |
|
61419 | 17 |
repeat the example from the previous section to illustrate this.\<close> |
27063 | 18 |
|
67398 | 19 |
interpretation %visible int: partial_order "(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool" |
61566
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents:
61565
diff
changeset
|
20 |
rewrites "int.less x y = (x < y)" |
32981 | 21 |
proof - |
67398 | 22 |
show "partial_order ((\<le>) :: int \<Rightarrow> int \<Rightarrow> bool)" |
32981 | 23 |
by unfold_locales auto |
67398 | 24 |
then interpret int: partial_order "(\<le>) :: [int, int] \<Rightarrow> bool" . |
43543
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
ballarin
parents:
37206
diff
changeset
|
25 |
show "int.less x y = (x < y)" |
32982 | 26 |
unfolding int.less_def by auto |
32981 | 27 |
qed |
27063 | 28 |
|
61419 | 29 |
text \<open>The inner interpretation is immediate from the preceding fact |
32981 | 30 |
and proved by assumption (Isar short hand ``.''). It enriches the |
31 |
local proof context by the theorems |
|
30826 | 32 |
also obtained in the interpretation from Section~\ref{sec:po-first}, |
69505 | 33 |
and \<open>int.less_def\<close> may directly be used to unfold the |
30826 | 34 |
definition. Theorems from the local interpretation disappear after |
32983 | 35 |
leaving the proof context --- that is, after the succeeding |
61419 | 36 |
\isakeyword{next} or \isakeyword{qed} statement.\<close> |
27063 | 37 |
|
38 |
||
61419 | 39 |
subsection \<open>Further Interpretations\<close> |
27063 | 40 |
|
61419 | 41 |
text \<open>Further interpretations are necessary for |
69505 | 42 |
the other locales. In \<open>lattice\<close> the operations~\<open>\<sqinter>\<close> |
69597 | 43 |
and~\<open>\<squnion>\<close> are substituted by \<^term>\<open>min :: int \<Rightarrow> int \<Rightarrow> int\<close> |
44 |
and \<^term>\<open>max :: int \<Rightarrow> int \<Rightarrow> int\<close>. The entire proof for the |
|
32981 | 45 |
interpretation is reproduced to give an example of a more |
32983 | 46 |
elaborate interpretation proof. Note that the equations are named |
61419 | 47 |
so they can be used in a later example.\<close> |
27063 | 48 |
|
67398 | 49 |
interpretation %visible int: lattice "(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool" |
61566
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents:
61565
diff
changeset
|
50 |
rewrites int_min_eq: "int.meet x y = min x y" |
43543
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
ballarin
parents:
37206
diff
changeset
|
51 |
and int_max_eq: "int.join x y = max x y" |
32981 | 52 |
proof - |
67398 | 53 |
show "lattice ((\<le>) :: int \<Rightarrow> int \<Rightarrow> bool)" |
61419 | 54 |
txt \<open>\normalsize We have already shown that this is a partial |
55 |
order,\<close> |
|
32981 | 56 |
apply unfold_locales |
61419 | 57 |
txt \<open>\normalsize hence only the lattice axioms remain to be |
43708 | 58 |
shown. |
32983 | 59 |
@{subgoals [display]} |
69505 | 60 |
By \<open>is_inf\<close> and \<open>is_sup\<close>,\<close> |
32982 | 61 |
apply (unfold int.is_inf_def int.is_sup_def) |
61419 | 62 |
txt \<open>\normalsize the goals are transformed to these |
43708 | 63 |
statements: |
64 |
@{subgoals [display]} |
|
65 |
This is Presburger arithmetic, which can be solved by the |
|
69505 | 66 |
method \<open>arith\<close>.\<close> |
32981 | 67 |
by arith+ |
61419 | 68 |
txt \<open>\normalsize In order to show the equations, we put ourselves |
32981 | 69 |
in a situation where the lattice theorems can be used in a |
61419 | 70 |
convenient way.\<close> |
67398 | 71 |
then interpret int: lattice "(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool" . |
43543
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
ballarin
parents:
37206
diff
changeset
|
72 |
show "int.meet x y = min x y" |
32982 | 73 |
by (bestsimp simp: int.meet_def int.is_inf_def) |
43543
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
ballarin
parents:
37206
diff
changeset
|
74 |
show "int.join x y = max x y" |
32982 | 75 |
by (bestsimp simp: int.join_def int.is_sup_def) |
32981 | 76 |
qed |
27063 | 77 |
|
69505 | 78 |
text \<open>Next follows that \<open>(\<le>)\<close> is a total order, again for |
61419 | 79 |
the integers.\<close> |
27063 | 80 |
|
67398 | 81 |
interpretation %visible int: total_order "(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool" |
32981 | 82 |
by unfold_locales arith |
27063 | 83 |
|
61419 | 84 |
text \<open>Theorems that are available in the theory at this point are shown in |
32982 | 85 |
Table~\ref{tab:int-lattice}. Two points are worth noting: |
27063 | 86 |
|
87 |
\begin{table} |
|
88 |
\hrule |
|
89 |
\vspace{2ex} |
|
90 |
\begin{center} |
|
91 |
\begin{tabular}{l} |
|
69505 | 92 |
@{thm [source] int.less_def} from locale \<open>partial_order\<close>: \\ |
32982 | 93 |
\quad @{thm int.less_def} \\ |
69505 | 94 |
@{thm [source] int.meet_left} from locale \<open>lattice\<close>: \\ |
63724 | 95 |
\quad @{thm int.meet_left} \\ |
69505 | 96 |
@{thm [source] int.join_distr} from locale \<open>distrib_lattice\<close>: \\ |
32982 | 97 |
\quad @{thm int.join_distr} \\ |
69505 | 98 |
@{thm [source] int.less_total} from locale \<open>total_order\<close>: \\ |
32982 | 99 |
\quad @{thm int.less_total} |
27063 | 100 |
\end{tabular} |
101 |
\end{center} |
|
102 |
\hrule |
|
69505 | 103 |
\caption{Interpreted theorems for~\<open>\<le>\<close> on the integers.} |
32982 | 104 |
\label{tab:int-lattice} |
27063 | 105 |
\end{table} |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents:
31960
diff
changeset
|
106 |
|
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents:
31960
diff
changeset
|
107 |
\begin{itemize} |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents:
31960
diff
changeset
|
108 |
\item |
69505 | 109 |
Locale \<open>distrib_lattice\<close> was also interpreted. Since the |
32983 | 110 |
locale hierarchy reflects that total orders are distributive |
32981 | 111 |
lattices, the interpretation of the latter was inserted |
112 |
automatically with the interpretation of the former. In general, |
|
32983 | 113 |
interpretation traverses the locale hierarchy upwards and interprets |
114 |
all encountered locales, regardless whether imported or proved via |
|
115 |
the \isakeyword{sublocale} command. Existing interpretations are |
|
116 |
skipped avoiding duplicate work. |
|
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents:
31960
diff
changeset
|
117 |
\item |
69597 | 118 |
The predicate \<^term>\<open>(<)\<close> appears in theorem @{thm [source] |
32983 | 119 |
int.less_total} |
69505 | 120 |
although an equation for the replacement of \<open>(\<sqsubset>)\<close> was only |
121 |
given in the interpretation of \<open>partial_order\<close>. The |
|
32983 | 122 |
interpretation equations are pushed downwards the hierarchy for |
123 |
related interpretations --- that is, for interpretations that share |
|
124 |
the instances of parameters they have in common. |
|
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents:
31960
diff
changeset
|
125 |
\end{itemize} |
61419 | 126 |
\<close> |
27063 | 127 |
|
61419 | 128 |
text \<open>The interpretations for a locale $n$ within the current |
32981 | 129 |
theory may be inspected with \isakeyword{print\_interps}~$n$. This |
130 |
prints the list of instances of $n$, for which interpretations exist. |
|
69597 | 131 |
For example, \isakeyword{print\_interps} \<^term>\<open>partial_order\<close> |
32981 | 132 |
outputs the following: |
32983 | 133 |
\begin{small} |
32981 | 134 |
\begin{alltt} |
67398 | 135 |
int : partial_order "(\(\le\))" |
32981 | 136 |
\end{alltt} |
32983 | 137 |
\end{small} |
138 |
Of course, there is only one interpretation. |
|
61731 | 139 |
The interpretation qualifier on the left is mandatory. Qualifiers |
140 |
can either be \emph{mandatory} or \emph{optional}. Optional qualifiers |
|
141 |
are designated by ``?''. Mandatory qualifiers must occur in |
|
142 |
name references while optional ones need not. Mandatory qualifiers |
|
32983 | 143 |
prevent accidental hiding of names, while optional qualifiers can be |
61731 | 144 |
more convenient to use. The default is mandatory. |
61419 | 145 |
\<close> |
27063 | 146 |
|
147 |
||
61419 | 148 |
section \<open>Locale Expressions \label{sec:expressions}\<close> |
27063 | 149 |
|
61419 | 150 |
text \<open> |
69597 | 151 |
A map~\<^term>\<open>\<phi>\<close> between partial orders~\<open>\<sqsubseteq>\<close> and~\<open>\<preceq>\<close> |
69505 | 152 |
is called order preserving if \<open>x \<sqsubseteq> y\<close> implies \<open>\<phi> x \<preceq> |
153 |
\<phi> y\<close>. This situation is more complex than those encountered so |
|
27063 | 154 |
far: it involves two partial orders, and it is desirable to use the |
155 |
existing locale for both. |
|
156 |
||
69505 | 157 |
A locale for order preserving maps requires three parameters: \<open>le\<close>~(\isakeyword{infixl}~\<open>\<sqsubseteq>\<close>) and \<open>le'\<close>~(\isakeyword{infixl}~\<open>\<preceq>\<close>) for the orders and~\<open>\<phi>\<close> |
32983 | 158 |
for the map. |
32981 | 159 |
|
160 |
In order to reuse the existing locale for partial orders, which has |
|
69505 | 161 |
the single parameter~\<open>le\<close>, it must be imported twice, once |
162 |
mapping its parameter to~\<open>le\<close> from the new locale and once |
|
163 |
to~\<open>le'\<close>. This can be achieved with a compound locale |
|
32981 | 164 |
expression. |
165 |
||
32983 | 166 |
In general, a locale expression is a sequence of \emph{locale instances} |
167 |
separated by~``$\textbf{+}$'' and followed by a \isakeyword{for} |
|
32981 | 168 |
clause. |
32983 | 169 |
An instance has the following format: |
32981 | 170 |
\begin{quote} |
171 |
\textit{qualifier} \textbf{:} \textit{locale-name} |
|
172 |
\textit{parameter-instantiation} |
|
173 |
\end{quote} |
|
174 |
We have already seen locale instances as arguments to |
|
175 |
\isakeyword{interpretation} in Section~\ref{sec:interpretation}. |
|
32983 | 176 |
As before, the qualifier serves to disambiguate names from |
61731 | 177 |
different instances of the same locale, and unless designated with a |
178 |
``?'' it must occur in name references. |
|
30750
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
179 |
|
69505 | 180 |
Since the parameters~\<open>le\<close> and~\<open>le'\<close> are to be partial |
32981 | 181 |
orders, our locale for order preserving maps will import the these |
182 |
instances: |
|
32983 | 183 |
\begin{small} |
32981 | 184 |
\begin{alltt} |
185 |
le: partial_order le |
|
186 |
le': partial_order le' |
|
187 |
\end{alltt} |
|
32983 | 188 |
\end{small} |
189 |
For matter of convenience we choose to name parameter names and |
|
190 |
qualifiers alike. This is an arbitrary decision. Technically, qualifiers |
|
32981 | 191 |
and parameters are unrelated. |
192 |
||
193 |
Having determined the instances, let us turn to the \isakeyword{for} |
|
194 |
clause. It serves to declare locale parameters in the same way as |
|
195 |
the context element \isakeyword{fixes} does. Context elements can |
|
196 |
only occur after the import section, and therefore the parameters |
|
32983 | 197 |
referred to in the instances must be declared in the \isakeyword{for} |
198 |
clause. The \isakeyword{for} clause is also where the syntax of these |
|
32981 | 199 |
parameters is declared. |
200 |
||
69505 | 201 |
Two context elements for the map parameter~\<open>\<phi>\<close> and the |
32983 | 202 |
assumptions that it is order preserving complete the locale |
61419 | 203 |
declaration.\<close> |
27063 | 204 |
|
205 |
locale order_preserving = |
|
30750
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
206 |
le: partial_order le + le': partial_order le' |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
76987
diff
changeset
|
207 |
for le (infixl \<open>\<sqsubseteq>\<close> 50) and le' (infixl \<open>\<preceq>\<close> 50) + |
32981 | 208 |
fixes \<phi> |
27063 | 209 |
assumes hom_le: "x \<sqsubseteq> y \<Longrightarrow> \<phi> x \<preceq> \<phi> y" |
210 |
||
61419 | 211 |
text (in order_preserving) \<open>Here are examples of theorems that are |
32981 | 212 |
available in the locale: |
27063 | 213 |
|
32983 | 214 |
\hspace*{1em}@{thm [source] hom_le}: @{thm hom_le} |
27063 | 215 |
|
32983 | 216 |
\hspace*{1em}@{thm [source] le.less_le_trans}: @{thm le.less_le_trans} |
27063 | 217 |
|
32983 | 218 |
\hspace*{1em}@{thm [source] le'.less_le_trans}: |
219 |
@{thm [display, indent=4] le'.less_le_trans} |
|
61701
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61566
diff
changeset
|
220 |
While there is infix syntax for the strict operation associated with |
69597 | 221 |
\<^term>\<open>(\<sqsubseteq>)\<close>, there is none for the strict version of \<^term>\<open>(\<preceq>)\<close>. The syntax \<open>\<sqsubset>\<close> for \<open>less\<close> is only |
32981 | 222 |
available for the original instance it was declared for. We may |
69505 | 223 |
introduce infix syntax for \<open>le'.less\<close> with the following declaration:\<close> |
27063 | 224 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
76987
diff
changeset
|
225 |
notation (in order_preserving) le'.less (infixl \<open>\<prec>\<close> 50) |
27063 | 226 |
|
61419 | 227 |
text (in order_preserving) \<open>Now the theorem is displayed nicely as |
30826 | 228 |
@{thm [source] le'.less_le_trans}: |
61419 | 229 |
@{thm [display, indent=2] le'.less_le_trans}\<close> |
27063 | 230 |
|
61419 | 231 |
text \<open>There are short notations for locale expressions. These are |
232 |
discussed in the following.\<close> |
|
32981 | 233 |
|
32983 | 234 |
|
61419 | 235 |
subsection \<open>Default Instantiations\<close> |
32981 | 236 |
|
61419 | 237 |
text \<open> |
32983 | 238 |
It is possible to omit parameter instantiations. The |
239 |
instantiation then defaults to the name of |
|
69505 | 240 |
the parameter itself. For example, the locale expression \<open>partial_order\<close> is short for \<open>partial_order le\<close>, since the |
241 |
locale's single parameter is~\<open>le\<close>. We took advantage of this |
|
32983 | 242 |
in the \isakeyword{sublocale} declarations of |
61419 | 243 |
Section~\ref{sec:changing-the-hierarchy}.\<close> |
32981 | 244 |
|
245 |
||
61419 | 246 |
subsection \<open>Implicit Parameters \label{sec:implicit-parameters}\<close> |
32983 | 247 |
|
61419 | 248 |
text \<open>In a locale expression that occurs within a locale |
32981 | 249 |
declaration, omitted parameters additionally extend the (possibly |
250 |
empty) \isakeyword{for} clause. |
|
27063 | 251 |
|
32983 | 252 |
The \isakeyword{for} clause is a general construct of Isabelle/Isar |
253 |
to mark names occurring in the preceding declaration as ``arbitrary |
|
254 |
but fixed''. This is necessary for example, if the name is already |
|
255 |
bound in a surrounding context. In a locale expression, names |
|
256 |
occurring in parameter instantiations should be bound by a |
|
257 |
\isakeyword{for} clause whenever these names are not introduced |
|
258 |
elsewhere in the context --- for example, on the left hand side of a |
|
259 |
\isakeyword{sublocale} declaration. |
|
27063 | 260 |
|
32981 | 261 |
There is an exception to this rule in locale declarations, where the |
32983 | 262 |
\isakeyword{for} clause serves to declare locale parameters. Here, |
32981 | 263 |
locale parameters for which no parameter instantiation is given are |
264 |
implicitly added, with their mixfix syntax, at the beginning of the |
|
265 |
\isakeyword{for} clause. For example, in a locale declaration, the |
|
69505 | 266 |
expression \<open>partial_order\<close> is short for |
32983 | 267 |
\begin{small} |
32981 | 268 |
\begin{alltt} |
269 |
partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.} |
|
270 |
\end{alltt} |
|
32983 | 271 |
\end{small} |
272 |
This short hand was used in the locale declarations throughout |
|
32981 | 273 |
Section~\ref{sec:import}. |
61419 | 274 |
\<close> |
32981 | 275 |
|
61419 | 276 |
text\<open> |
32983 | 277 |
The following locale declarations provide more examples. A |
69505 | 278 |
map~\<open>\<phi>\<close> is a lattice homomorphism if it preserves meet and |
61419 | 279 |
join.\<close> |
30750
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
280 |
|
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
281 |
locale lattice_hom = |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
76987
diff
changeset
|
282 |
le: lattice + le': lattice le' for le' (infixl \<open>\<preceq>\<close> 50) + |
27063 | 283 |
fixes \<phi> |
30826 | 284 |
assumes hom_meet: "\<phi> (x \<sqinter> y) = le'.meet (\<phi> x) (\<phi> y)" |
285 |
and hom_join: "\<phi> (x \<squnion> y) = le'.join (\<phi> x) (\<phi> y)" |
|
27063 | 286 |
|
69597 | 287 |
text \<open>The parameter instantiation in the first instance of \<^term>\<open>lattice\<close> is omitted. This causes the parameter~\<open>le\<close> to be |
32983 | 288 |
added to the \isakeyword{for} clause, and the locale has |
69505 | 289 |
parameters~\<open>le\<close>,~\<open>le'\<close> and, of course,~\<open>\<phi>\<close>. |
27063 | 290 |
|
32981 | 291 |
Before turning to the second example, we complete the locale by |
32983 | 292 |
providing infix syntax for the meet and join operations of the |
32981 | 293 |
second lattice. |
61419 | 294 |
\<close> |
32981 | 295 |
|
46855 | 296 |
context lattice_hom |
297 |
begin |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
76987
diff
changeset
|
298 |
notation le'.meet (infixl \<open>\<sqinter>''\<close> 50) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
76987
diff
changeset
|
299 |
notation le'.join (infixl \<open>\<squnion>''\<close> 50) |
32981 | 300 |
end |
301 |
||
61419 | 302 |
text \<open>The next example makes radical use of the short hand |
32983 | 303 |
facilities. A homomorphism is an endomorphism if both orders |
61419 | 304 |
coincide.\<close> |
27063 | 305 |
|
30750
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
306 |
locale lattice_end = lattice_hom _ le |
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
307 |
|
69505 | 308 |
text \<open>The notation~\<open>_\<close> enables to omit a parameter in a |
309 |
positional instantiation. The omitted parameter,~\<open>le\<close> becomes |
|
32981 | 310 |
the parameter of the declared locale and is, in the following |
69505 | 311 |
position, used to instantiate the second parameter of \<open>lattice_hom\<close>. The effect is that of identifying the first in second |
61419 | 312 |
parameter of the homomorphism locale.\<close> |
27063 | 313 |
|
61419 | 314 |
text \<open>The inheritance diagram of the situation we have now is shown |
27063 | 315 |
in Figure~\ref{fig:hom}, where the dashed line depicts an |
32983 | 316 |
interpretation which is introduced below. Parameter instantiations |
317 |
are indicated by $\sqsubseteq \mapsto \preceq$ etc. By looking at |
|
318 |
the inheritance diagram it would seem |
|
69505 | 319 |
that two identical copies of each of the locales \<open>partial_order\<close> and \<open>lattice\<close> are imported by \<open>lattice_end\<close>. This is not the case! Inheritance paths with |
32981 | 320 |
identical morphisms are automatically detected and |
27503 | 321 |
the conclusions of the respective locales appear only once. |
27063 | 322 |
|
323 |
\begin{figure} |
|
324 |
\hrule \vspace{2ex} |
|
325 |
\begin{center} |
|
326 |
\begin{tikzpicture} |
|
69505 | 327 |
\node (o) at (0,0) {\<open>partial_order\<close>}; |
328 |
\node (oh) at (1.5,-2) {\<open>order_preserving\<close>}; |
|
27063 | 329 |
\node (oh1) at (1.5,-0.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; |
330 |
\node (oh2) at (0,-1.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; |
|
69505 | 331 |
\node (l) at (-1.5,-2) {\<open>lattice\<close>}; |
332 |
\node (lh) at (0,-4) {\<open>lattice_hom\<close>}; |
|
27063 | 333 |
\node (lh1) at (0,-2.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; |
334 |
\node (lh2) at (-1.5,-3.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; |
|
69505 | 335 |
\node (le) at (0,-6) {\<open>lattice_end\<close>}; |
27063 | 336 |
\node (le1) at (0,-4.8) |
337 |
[anchor=west]{$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; |
|
338 |
\node (le2) at (0,-5.2) |
|
339 |
[anchor=west]{$\scriptscriptstyle \preceq \mapsto \sqsubseteq$}; |
|
340 |
\draw (o) -- (l); |
|
341 |
\draw[dashed] (oh) -- (lh); |
|
342 |
\draw (lh) -- (le); |
|
343 |
\draw (o) .. controls (oh1.south west) .. (oh); |
|
344 |
\draw (o) .. controls (oh2.north east) .. (oh); |
|
345 |
\draw (l) .. controls (lh1.south west) .. (lh); |
|
346 |
\draw (l) .. controls (lh2.north east) .. (lh); |
|
347 |
\end{tikzpicture} |
|
348 |
\end{center} |
|
349 |
\hrule |
|
350 |
\caption{Hierarchy of Homomorphism Locales.} |
|
351 |
\label{fig:hom} |
|
352 |
\end{figure} |
|
61419 | 353 |
\<close> |
27063 | 354 |
|
61419 | 355 |
text \<open>It can be shown easily that a lattice homomorphism is order |
27063 | 356 |
preserving. As the final example of this section, a locale |
61419 | 357 |
interpretation is used to assert this:\<close> |
27063 | 358 |
|
46855 | 359 |
sublocale lattice_hom \<subseteq> order_preserving |
360 |
proof unfold_locales |
|
27063 | 361 |
fix x y |
362 |
assume "x \<sqsubseteq> y" |
|
363 |
then have "y = (x \<squnion> y)" by (simp add: le.join_connection) |
|
364 |
then have "\<phi> y = (\<phi> x \<squnion>' \<phi> y)" by (simp add: hom_join [symmetric]) |
|
365 |
then show "\<phi> x \<preceq> \<phi> y" by (simp add: le'.join_connection) |
|
366 |
qed |
|
367 |
||
61419 | 368 |
text (in lattice_hom) \<open> |
30393
aa6f42252bf6
replaced old locale option by proper "text (in locale)";
wenzelm
parents:
29568
diff
changeset
|
369 |
Theorems and other declarations --- syntax, in particular --- from |
69505 | 370 |
the locale \<open>order_preserving\<close> are now active in \<open>lattice_hom\<close>, for example |
30826 | 371 |
@{thm [source] hom_le}: |
372 |
@{thm [display, indent=2] hom_le} |
|
32983 | 373 |
This theorem will be useful in the following section. |
61419 | 374 |
\<close> |
27063 | 375 |
|
32983 | 376 |
|
61419 | 377 |
section \<open>Conditional Interpretation\<close> |
27063 | 378 |
|
61419 | 379 |
text \<open>There are situations where an interpretation is not possible |
32983 | 380 |
in the general case since the desired property is only valid if |
381 |
certain conditions are fulfilled. Take, for example, the function |
|
69505 | 382 |
\<open>\<lambda>i. n * i\<close> that scales its argument by a constant factor. |
32983 | 383 |
This function is order preserving (and even a lattice endomorphism) |
69597 | 384 |
with respect to \<^term>\<open>(\<le>)\<close> provided \<open>n \<ge> 0\<close>. |
32983 | 385 |
|
386 |
It is not possible to express this using a global interpretation, |
|
69597 | 387 |
because it is in general unspecified whether~\<^term>\<open>n\<close> is |
32983 | 388 |
non-negative, but one may make an interpretation in an inner context |
389 |
of a proof where full information is available. |
|
390 |
This is not fully satisfactory either, since potentially |
|
391 |
interpretations may be required to make interpretations in many |
|
392 |
contexts. What is |
|
393 |
required is an interpretation that depends on the condition --- and |
|
394 |
this can be done with the \isakeyword{sublocale} command. For this |
|
61419 | 395 |
purpose, we introduce a locale for the condition.\<close> |
32983 | 396 |
|
32981 | 397 |
locale non_negative = |
398 |
fixes n :: int |
|
32983 | 399 |
assumes non_neg: "0 \<le> n" |
400 |
||
61419 | 401 |
text \<open>It is again convenient to make the interpretation in an |
63390 | 402 |
incremental fashion, first for order preserving maps, then for |
61419 | 403 |
lattice endomorphisms.\<close> |
32981 | 404 |
|
32983 | 405 |
sublocale non_negative \<subseteq> |
67398 | 406 |
order_preserving "(\<le>)" "(\<le>)" "\<lambda>i. n * i" |
32983 | 407 |
using non_neg by unfold_locales (rule mult_left_mono) |
408 |
||
61419 | 409 |
text \<open>While the proof of the previous interpretation |
69597 | 410 |
is straightforward from monotonicity lemmas for~\<^term>\<open>(*)\<close>, the |
61419 | 411 |
second proof follows a useful pattern.\<close> |
32981 | 412 |
|
67398 | 413 |
sublocale %visible non_negative \<subseteq> lattice_end "(\<le>)" "\<lambda>i. n * i" |
32983 | 414 |
proof (unfold_locales, unfold int_min_eq int_max_eq) |
61419 | 415 |
txt \<open>\normalsize Unfolding the locale predicates \emph{and} the |
32983 | 416 |
interpretation equations immediately yields two subgoals that |
417 |
reflect the core conjecture. |
|
418 |
@{subgoals [display]} |
|
69597 | 419 |
It is now necessary to show, in the context of \<^term>\<open>non_negative\<close>, that multiplication by~\<^term>\<open>n\<close> commutes with |
420 |
\<^term>\<open>min\<close> and \<^term>\<open>max\<close>.\<close> |
|
32983 | 421 |
qed (auto simp: hom_le) |
32981 | 422 |
|
61419 | 423 |
text (in order_preserving) \<open>The lemma @{thm [source] hom_le} |
32983 | 424 |
simplifies a proof that would have otherwise been lengthy and we may |
61419 | 425 |
consider making it a default rule for the simplifier:\<close> |
32981 | 426 |
|
32983 | 427 |
lemmas (in order_preserving) hom_le [simp] |
32981 | 428 |
|
429 |
||
61419 | 430 |
subsection \<open>Avoiding Infinite Chains of Interpretations |
431 |
\label{sec:infinite-chains}\<close> |
|
32983 | 432 |
|
61419 | 433 |
text \<open>Similar situations arise frequently in formalisations of |
32983 | 434 |
abstract algebra where it is desirable to express that certain |
435 |
constructions preserve certain properties. For example, polynomials |
|
436 |
over rings are rings, or --- an example from the domain where the |
|
437 |
illustrations of this tutorial are taken from --- a partial order |
|
438 |
may be obtained for a function space by point-wise lifting of the |
|
439 |
partial order of the co-domain. This corresponds to the following |
|
61419 | 440 |
interpretation:\<close> |
32983 | 441 |
|
442 |
sublocale %visible partial_order \<subseteq> f: partial_order "\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x" |
|
443 |
oops |
|
444 |
||
61419 | 445 |
text \<open>Unfortunately this is a cyclic interpretation that leads to an |
32983 | 446 |
infinite chain, namely |
447 |
@{text [display, indent=2] "partial_order \<subseteq> partial_order (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) \<subseteq> |
|
448 |
partial_order (\<lambda>f g. \<forall>x y. f x y \<sqsubseteq> g x y) \<subseteq> \<dots>"} |
|
449 |
and the interpretation is rejected. |
|
450 |
||
451 |
Instead it is necessary to declare a locale that is logically |
|
69597 | 452 |
equivalent to \<^term>\<open>partial_order\<close> but serves to collect facts |
32983 | 453 |
about functions spaces where the co-domain is a partial order, and |
61419 | 454 |
to make the interpretation in its context:\<close> |
32981 | 455 |
|
32983 | 456 |
locale fun_partial_order = partial_order |
457 |
||
458 |
sublocale fun_partial_order \<subseteq> |
|
459 |
f: partial_order "\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x" |
|
460 |
by unfold_locales (fast,rule,fast,blast intro: trans) |
|
461 |
||
61419 | 462 |
text \<open>It is quite common in abstract algebra that such a construction |
32983 | 463 |
maps a hierarchy of algebraic structures (or specifications) to a |
464 |
related hierarchy. By means of the same lifting, a function space |
|
61419 | 465 |
is a lattice if its co-domain is a lattice:\<close> |
32983 | 466 |
|
467 |
locale fun_lattice = fun_partial_order + lattice |
|
32981 | 468 |
|
32983 | 469 |
sublocale fun_lattice \<subseteq> f: lattice "\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x" |
46855 | 470 |
proof unfold_locales |
32983 | 471 |
fix f g |
472 |
have "partial_order.is_inf (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g (\<lambda>x. f x \<sqinter> g x)" |
|
61565
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents:
61419
diff
changeset
|
473 |
apply (rule f.is_infI) apply rule+ apply (drule spec, assumption)+ done |
32983 | 474 |
then show "\<exists>inf. partial_order.is_inf (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g inf" |
475 |
by fast |
|
476 |
next |
|
477 |
fix f g |
|
478 |
have "partial_order.is_sup (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g (\<lambda>x. f x \<squnion> g x)" |
|
61565
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents:
61419
diff
changeset
|
479 |
apply (rule f.is_supI) apply rule+ apply (drule spec, assumption)+ done |
32983 | 480 |
then show "\<exists>sup. partial_order.is_sup (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g sup" |
481 |
by fast |
|
482 |
qed |
|
483 |
||
27063 | 484 |
|
61419 | 485 |
section \<open>Further Reading\<close> |
27063 | 486 |
|
61419 | 487 |
text \<open>More information on locales and their interpretation is |
27063 | 488 |
available. For the locale hierarchy of import and interpretation |
76987 | 489 |
dependencies see~\<^cite>\<open>Ballarin2006a\<close>; interpretations in theories |
490 |
and proofs are covered in~\<^cite>\<open>Ballarin2006b\<close>. In the latter, I |
|
27063 | 491 |
show how interpretation in proofs enables to reason about families |
492 |
of algebraic structures, which cannot be expressed with locales |
|
493 |
directly. |
|
494 |
||
76987 | 495 |
Haftmann and Wenzel~\<^cite>\<open>HaftmannWenzel2007\<close> overcome a restriction |
27063 | 496 |
of axiomatic type classes through a combination with locale |
497 |
interpretation. The result is a Haskell-style class system with a |
|
30750
3779e2158dad
Update explanation of locale expressions to locale reimplementation.
ballarin
parents:
30580
diff
changeset
|
498 |
facility to generate ML and Haskell code. Classes are sufficient for |
27063 | 499 |
simple specifications with a single type parameter. The locales for |
500 |
orders and lattices presented in this tutorial fall into this |
|
501 |
category. Order preserving maps, homomorphisms and vector spaces, |
|
502 |
on the other hand, do not. |
|
503 |
||
32981 | 504 |
The locales reimplementation for Isabelle 2009 provides, among other |
32983 | 505 |
improvements, a clean integration with Isabelle/Isar's local theory |
506 |
mechanisms, which are described in another paper by Haftmann and |
|
76987 | 507 |
Wenzel~\<^cite>\<open>HaftmannWenzel2009\<close>. |
32981 | 508 |
|
76987 | 509 |
The original work of Kammüller on locales~\<^cite>\<open>KammullerEtAl1999\<close> |
32983 | 510 |
may be of interest from a historical perspective. My previous |
76987 | 511 |
report on locales and locale expressions~\<^cite>\<open>Ballarin2004a\<close> |
32983 | 512 |
describes a simpler form of expressions than available now and is |
513 |
outdated. The mathematical background on orders and lattices is |
|
76987 | 514 |
taken from Jacobson's textbook on algebra~\<^cite>\<open>\<open>Chapter~8\<close> in Jacobson1985\<close>. |
32981 | 515 |
|
516 |
The sources of this tutorial, which include all proofs, are |
|
517 |
available with the Isabelle distribution at |
|
68224 | 518 |
\<^url>\<open>https://isabelle.in.tum.de\<close>. |
61419 | 519 |
\<close> |
27063 | 520 |
|
61419 | 521 |
text \<open> |
27063 | 522 |
\begin{table} |
523 |
\hrule |
|
524 |
\vspace{2ex} |
|
525 |
\begin{center} |
|
526 |
\begin{tabular}{l>$c<$l} |
|
527 |
\multicolumn{3}{l}{Miscellaneous} \\ |
|
528 |
||
529 |
\textit{attr-name} & ::= |
|
530 |
& \textit{name} $|$ \textit{attribute} $|$ |
|
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
531 |
\textit{name} \textit{attribute} \\ |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
532 |
\textit{qualifier} & ::= |
61731 | 533 |
& \textit{name} [``\textbf{?}''] \\[2ex] |
27063 | 534 |
|
535 |
\multicolumn{3}{l}{Context Elements} \\ |
|
536 |
||
537 |
\textit{fixes} & ::= |
|
538 |
& \textit{name} [ ``\textbf{::}'' \textit{type} ] |
|
539 |
[ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ |
|
540 |
\textit{mixfix} ] \\ |
|
541 |
\begin{comment} |
|
542 |
\textit{constrains} & ::= |
|
543 |
& \textit{name} ``\textbf{::}'' \textit{type} \\ |
|
544 |
\end{comment} |
|
545 |
\textit{assumes} & ::= |
|
546 |
& [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ |
|
547 |
\begin{comment} |
|
548 |
\textit{defines} & ::= |
|
549 |
& [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ |
|
550 |
\textit{notes} & ::= |
|
551 |
& [ \textit{attr-name} ``\textbf{=}'' ] |
|
552 |
( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ |
|
553 |
\end{comment} |
|
554 |
||
555 |
\textit{element} & ::= |
|
556 |
& \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ |
|
557 |
\begin{comment} |
|
558 |
& | |
|
559 |
& \textbf{constrains} \textit{constrains} |
|
560 |
( \textbf{and} \textit{constrains} )$^*$ \\ |
|
561 |
\end{comment} |
|
562 |
& | |
|
563 |
& \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\[2ex] |
|
564 |
%\begin{comment} |
|
565 |
% & | |
|
566 |
% & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ |
|
567 |
% & | |
|
568 |
% & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ |
|
569 |
%\end{comment} |
|
570 |
||
571 |
\multicolumn{3}{l}{Locale Expressions} \\ |
|
572 |
||
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
573 |
\textit{pos-insts} & ::= |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
574 |
& ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\ |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
575 |
\textit{named-insts} & ::= |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
576 |
& \textbf{where} \textit{name} ``\textbf{=}'' \textit{term} |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
577 |
( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\ |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
578 |
\textit{instance} & ::= |
30751 | 579 |
& [ \textit{qualifier} ``\textbf{:}'' ] |
33867 | 580 |
\textit{name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\ |
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
581 |
\textit{expression} & ::= |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
582 |
& \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$ |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
583 |
[ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex] |
27063 | 584 |
|
585 |
\multicolumn{3}{l}{Declaration of Locales} \\ |
|
586 |
||
587 |
\textit{locale} & ::= |
|
588 |
& \textit{element}$^+$ \\ |
|
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
589 |
& | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ |
27063 | 590 |
\textit{toplevel} & ::= |
591 |
& \textbf{locale} \textit{name} [ ``\textbf{=}'' |
|
592 |
\textit{locale} ] \\[2ex] |
|
593 |
||
594 |
\multicolumn{3}{l}{Interpretation} \\ |
|
595 |
||
596 |
\textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] |
|
597 |
\textit{prop} \\ |
|
61566
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents:
61565
diff
changeset
|
598 |
\textit{equations} & ::= & \textbf{rewrites} \textit{equation} ( \textbf{and} |
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
599 |
\textit{equation} )$^*$ \\ |
27063 | 600 |
\textit{toplevel} & ::= |
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
601 |
& \textbf{sublocale} \textit{name} ( ``$<$'' $|$ |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
602 |
``$\subseteq$'' ) \textit{expression} \textit{proof} \\ |
27063 | 603 |
& | |
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
604 |
& \textbf{interpretation} |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
605 |
\textit{expression} [ \textit{equations} ] \textit{proof} \\ |
27063 | 606 |
& | |
29566
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
607 |
& \textbf{interpret} |
937baa077df2
Fixed tutorial to compile with new locales; grammar of new locale commands.
ballarin
parents:
27595
diff
changeset
|
608 |
\textit{expression} \textit{proof} \\[2ex] |
27063 | 609 |
|
610 |
\multicolumn{3}{l}{Diagnostics} \\ |
|
611 |
||
612 |
\textit{toplevel} & ::= |
|
32981 | 613 |
& \textbf{print\_locales} \\ |
33867 | 614 |
& | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{name} \\ |
615 |
& | & \textbf{print\_interps} \textit{name} |
|
27063 | 616 |
\end{tabular} |
617 |
\end{center} |
|
618 |
\hrule |
|
30826 | 619 |
\caption{Syntax of Locale Commands.} |
27063 | 620 |
\label{tab:commands} |
621 |
\end{table} |
|
61419 | 622 |
\<close> |
27063 | 623 |
|
61732
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
624 |
text \<open>\textbf{Revision History.} The present fourth revision of the |
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
625 |
tutorial accommodates minor updates to the syntax of the locale commands |
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
626 |
and the handling of notation under morphisms introduced with Isabelle 2016. |
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
627 |
For the third revision of the tutorial, which corresponds to the published |
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
628 |
version, much of the explanatory text was rewritten. Inheritance of |
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
629 |
interpretation equations was made available with Isabelle 2009-1. |
32983 | 630 |
The second revision accommodates changes introduced by the locales |
631 |
reimplementation for Isabelle 2009. Most notably locale expressions |
|
61732
4653d0835e6e
Updates to the revision history of the locales tutorial.
ballarin
parents:
61731
diff
changeset
|
632 |
were generalised from renaming to instantiation.\<close> |
32981 | 633 |
|
61419 | 634 |
text \<open>\textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, |
37206 | 635 |
Randy Pollack, Andreas Schropp, Christian Sternagel and Makarius Wenzel |
636 |
have made |
|
32983 | 637 |
useful comments on earlier versions of this document. The section |
638 |
on conditional interpretation was inspired by a number of e-mail |
|
639 |
enquiries the author received from locale users, and which suggested |
|
640 |
that this use case is important enough to deserve explicit |
|
641 |
explanation. The term \emph{conditional interpretation} is due to |
|
61419 | 642 |
Larry Paulson.\<close> |
27063 | 643 |
|
644 |
end |