# HG changeset patch # User wenzelm # Date 1346095156 -7200 # Node ID 54da920baf383d273c4764c3f849b8cc0c835e3a # Parent 75d8778f94d31370e424c9feec1be29688809fac more standard document preparation within session context; diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Examples.thy Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,770 @@ +theory Examples +imports Main +begin + +pretty_setmargin %invisible 65 + +(* +text {* The following presentation will use notation of + Isabelle's meta logic, hence a few sentences to explain this. + The logical + primitives are universal quantification (@{text "\"}), entailment + (@{text "\"}) and equality (@{text "\"}). Variables (not bound + variables) are sometimes preceded by a question mark. The logic is + typed. Type variables are denoted by~@{text "'a"},~@{text "'b"} + etc., and~@{text "\"} is the function type. Double brackets~@{text + "\"} and~@{text "\"} are used to abbreviate nested entailment. +*} +*) + +section {* Introduction *} + +text {* + Locales are based on contexts. A \emph{context} can be seen as a + formula schema +\[ + @{text "\x\<^sub>1\x\<^sub>n. \ A\<^sub>1; \ ;A\<^sub>m \ \ \"} +\] + where the variables~@{text "x\<^sub>1"}, \ldots,~@{text "x\<^sub>n"} are called + \emph{parameters} and the premises $@{text "A\<^sub>1"}, \ldots,~@{text + "A\<^sub>m"}$ \emph{assumptions}. A formula~@{text "C"} + is a \emph{theorem} in the context if it is a conclusion +\[ + @{text "\x\<^sub>1\x\<^sub>n. \ A\<^sub>1; \ ;A\<^sub>m \ \ C"}. +\] + Isabelle/Isar's notion of context goes beyond this logical view. + Its contexts record, in a consecutive order, proved + conclusions along with \emph{attributes}, which can provide context + specific configuration information for proof procedures and concrete + syntax. From a logical perspective, locales are just contexts that + have been made persistent. To the user, though, they provide + powerful means for declaring and combining contexts, and for the + reuse of theorems proved in these contexts. + *} + +section {* Simple Locales *} + +text {* + In its simplest form, a + \emph{locale declaration} consists of a sequence of context elements + declaring parameters (keyword \isakeyword{fixes}) and assumptions + (keyword \isakeyword{assumes}). The following is the specification of + partial orders, as locale @{text partial_order}. + *} + + locale partial_order = + fixes le :: "'a \ 'a \ bool" (infixl "\" 50) + assumes refl [intro, simp]: "x \ x" + and anti_sym [intro]: "\ x \ y; y \ x \ \ x = y" + and trans [trans]: "\ x \ y; y \ z \ \ x \ z" + +text (in partial_order) {* The parameter of this locale is~@{text le}, + which is a binary predicate with infix syntax~@{text \}. The + parameter syntax is available in the subsequent + assumptions, which are the familiar partial order axioms. + + Isabelle recognises unbound names as free variables. In locale + assumptions, these are implicitly universally quantified. That is, + @{term "\ x \ y; y \ z \ \ x \ z"} in fact means + @{term "\x y z. \ x \ y; y \ z \ \ x \ z"}. + + Two commands are provided to inspect locales: + \isakeyword{print\_locales} lists the names of all locales of the + current theory; \isakeyword{print\_locale}~$n$ prints the parameters + and assumptions of locale $n$; the variation \isakeyword{print\_locale!}~$n$ + additionally outputs the conclusions that are stored in the locale. + We may inspect the new locale + by issuing \isakeyword{print\_locale!} @{term partial_order}. The output + is the following list of context elements. +\begin{small} +\begin{alltt} + \isakeyword{fixes} le :: "'a \(\Rightarrow\) 'a \(\Rightarrow\) bool" (\isakeyword{infixl} "\(\sqsubseteq\)" 50) + \isakeyword{assumes} "partial_order op \(\sqsubseteq\)" + \isakeyword{notes} assumption + refl [intro, simp] = `?x \(\sqsubseteq\) ?x` + \isakeyword{and} + anti_sym [intro] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?x\(\isasymrbrakk\) \(\Longrightarrow\) ?x = ?y` + \isakeyword{and} + trans [trans] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?z\(\isasymrbrakk\) \(\Longrightarrow\) ?x \(\sqsubseteq\) ?z` +\end{alltt} +\end{small} + The keyword \isakeyword{notes} denotes a conclusion element. There + is one conclusion, which was added automatically. Instead, there is + only one assumption, namely @{term "partial_order le"}. The locale + declaration has introduced the predicate @{term + partial_order} to the theory. This predicate is the + \emph{locale predicate}. Its definition may be inspected by + issuing \isakeyword{thm} @{thm [source] partial_order_def}. + @{thm [display, indent=2] partial_order_def} + In our example, this is a unary predicate over the parameter of the + locale. It is equivalent to the original assumptions, which have + been turned into conclusions and are + available as theorems in the context of the locale. The names and + attributes from the locale declaration are associated to these + theorems and are effective in the context of the locale. + + Each conclusion has a \emph{foundational theorem} as counterpart + in the theory. Technically, this is simply the theorem composed + of context and conclusion. For the transitivity theorem, this is + @{thm [source] partial_order.trans}: + @{thm [display, indent=2] partial_order.trans} +*} + +subsection {* Targets: Extending Locales *} + +text {* + The specification of a locale is fixed, but its list of conclusions + may be extended through Isar commands that take a \emph{target} argument. + In the following, \isakeyword{definition} and + \isakeyword{theorem} are illustrated. + Table~\ref{tab:commands-with-target} lists Isar commands that accept + a target. Isar provides various ways of specifying the target. A + target for a single command may be indicated with keyword + \isakeyword{in} in the following way: + +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{ll} + \isakeyword{definition} & definition through an equation \\ + \isakeyword{inductive} & inductive definition \\ + \isakeyword{primrec} & primitive recursion \\ + \isakeyword{fun}, \isakeyword{function} & general recursion \\ + \isakeyword{abbreviation} & syntactic abbreviation \\ + \isakeyword{theorem}, etc.\ & theorem statement with proof \\ + \isakeyword{theorems}, etc.\ & redeclaration of theorems \\ + \isakeyword{text}, etc.\ & document markup +\end{tabular} +\end{center} +\hrule +\caption{Isar commands that accept a target.} +\label{tab:commands-with-target} +\end{table} + *} + + definition (in partial_order) + less :: "'a \ 'a \ bool" (infixl "\" 50) + where "(x \ y) = (x \ y \ x \ y)" + +text (in partial_order) {* The strict order @{text less} with infix + syntax~@{text \} is + defined in terms of the locale parameter~@{text le} and the general + equality of the object logic we work in. The definition generates a + \emph{foundational constant} + @{term partial_order.less} with definition @{thm [source] + partial_order.less_def}: + @{thm [display, indent=2] partial_order.less_def} + At the same time, the locale is extended by syntax transformations + hiding this construction in the context of the locale. Here, the + abbreviation @{text less} is available for + @{text "partial_order.less le"}, and it is printed + and parsed as infix~@{text \}. Finally, the conclusion @{thm [source] + less_def} is added to the locale: + @{thm [display, indent=2] less_def} +*} + +text {* The treatment of theorem statements is more straightforward. + As an example, here is the derivation of a transitivity law for the + strict order relation. *} + + lemma (in partial_order) less_le_trans [trans]: + "\ x \ y; y \ z \ \ x \ z" + unfolding %visible less_def by %visible (blast intro: trans) + +text {* In the context of the proof, conclusions of the + locale may be used like theorems. Attributes are effective: @{text + anti_sym} was + declared as introduction rule, hence it is in the context's set of + rules used by the classical reasoner by default. *} + +subsection {* Context Blocks *} + +text {* When working with locales, sequences of commands with the same + target are frequent. A block of commands, delimited by + \isakeyword{begin} and \isakeyword{end}, makes a theory-like style + of working possible. All commands inside the block refer to the + same target. A block may immediately follow a locale + declaration, which makes that locale the target. Alternatively the + target for a block may be given with the \isakeyword{context} + command. + + This style of working is illustrated in the block below, where + notions of infimum and supremum for partial orders are introduced, + together with theorems about their uniqueness. *} + + context partial_order + begin + + definition + is_inf where "is_inf x y i = + (i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i))" + + definition + is_sup where "is_sup x y s = + (x \ s \ y \ s \ (\z. x \ z \ y \ z \ s \ z))" + + lemma %invisible is_infI [intro?]: "i \ x \ i \ y \ + (\z. z \ x \ z \ y \ z \ i) \ is_inf x y i" + by (unfold is_inf_def) blast + + lemma %invisible is_inf_lower [elim?]: + "is_inf x y i \ (i \ x \ i \ y \ C) \ C" + by (unfold is_inf_def) blast + + lemma %invisible is_inf_greatest [elim?]: + "is_inf x y i \ z \ x \ z \ y \ z \ i" + by (unfold is_inf_def) blast + + theorem is_inf_uniq: "\is_inf x y i; is_inf x y i'\ \ i = i'" + proof - + assume inf: "is_inf x y i" + assume inf': "is_inf x y i'" + show ?thesis + proof (rule anti_sym) + from inf' show "i \ i'" + proof (rule is_inf_greatest) + from inf show "i \ x" .. + from inf show "i \ y" .. + qed + from inf show "i' \ i" + proof (rule is_inf_greatest) + from inf' show "i' \ x" .. + from inf' show "i' \ y" .. + qed + qed + qed + + theorem %invisible is_inf_related [elim?]: "x \ y \ is_inf x y x" + proof - + assume "x \ y" + show ?thesis + proof + show "x \ x" .. + show "x \ y" by fact + fix z assume "z \ x" and "z \ y" show "z \ x" by fact + qed + qed + + lemma %invisible is_supI [intro?]: "x \ s \ y \ s \ + (\z. x \ z \ y \ z \ s \ z) \ is_sup x y s" + by (unfold is_sup_def) blast + + lemma %invisible is_sup_least [elim?]: + "is_sup x y s \ x \ z \ y \ z \ s \ z" + by (unfold is_sup_def) blast + + lemma %invisible is_sup_upper [elim?]: + "is_sup x y s \ (x \ s \ y \ s \ C) \ C" + by (unfold is_sup_def) blast + + theorem is_sup_uniq: "\is_sup x y s; is_sup x y s'\ \ s = s'" + proof - + assume sup: "is_sup x y s" + assume sup': "is_sup x y s'" + show ?thesis + proof (rule anti_sym) + from sup show "s \ s'" + proof (rule is_sup_least) + from sup' show "x \ s'" .. + from sup' show "y \ s'" .. + qed + from sup' show "s' \ s" + proof (rule is_sup_least) + from sup show "x \ s" .. + from sup show "y \ s" .. + qed + qed + qed + + theorem %invisible is_sup_related [elim?]: "x \ y \ is_sup x y y" + proof - + assume "x \ y" + show ?thesis + proof + show "x \ y" by fact + show "y \ y" .. + fix z assume "x \ z" and "y \ z" + show "y \ z" by fact + qed + qed + + end + +text {* The syntax of the locale commands discussed in this tutorial is + shown in Table~\ref{tab:commands}. The grammar is complete with the + exception of the context elements \isakeyword{constrains} and + \isakeyword{defines}, which are provided for backward + compatibility. See the Isabelle/Isar Reference + Manual~\cite{IsarRef} for full documentation. *} + + +section {* Import \label{sec:import} *} + +text {* + Algebraic structures are commonly defined by adding operations and + properties to existing structures. For example, partial orders + are extended to lattices and total orders. Lattices are extended to + distributive lattices. *} + +text {* + With locales, this kind of inheritance is achieved through + \emph{import} of locales. The import part of a locale declaration, + if present, precedes the context elements. Here is an example, + where partial orders are extended to lattices. + *} + + locale lattice = partial_order + + assumes ex_inf: "\inf. is_inf x y inf" + and ex_sup: "\sup. is_sup x y sup" + begin + +text {* These assumptions refer to the predicates for infimum + and supremum defined for @{text partial_order} in the previous + section. We now introduce the notions of meet and join. *} + + definition + meet (infixl "\" 70) where "x \ y = (THE inf. is_inf x y inf)" + definition + join (infixl "\" 65) where "x \ y = (THE sup. is_sup x y sup)" + + lemma %invisible meet_equality [elim?]: "is_inf x y i \ x \ y = i" + proof (unfold meet_def) + assume "is_inf x y i" + then show "(THE i. is_inf x y i) = i" + by (rule the_equality) (rule is_inf_uniq [OF _ `is_inf x y i`]) + qed + + lemma %invisible meetI [intro?]: + "i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i) \ x \ y = i" + by (rule meet_equality, rule is_infI) blast+ + + lemma %invisible is_inf_meet [intro?]: "is_inf x y (x \ y)" + proof (unfold meet_def) + from ex_inf obtain i where "is_inf x y i" .. + then show "is_inf x y (THE i. is_inf x y i)" + by (rule theI) (rule is_inf_uniq [OF _ `is_inf x y i`]) + qed + + lemma %invisible meet_left [intro?]: + "x \ y \ x" + by (rule is_inf_lower) (rule is_inf_meet) + + lemma %invisible meet_right [intro?]: + "x \ y \ y" + by (rule is_inf_lower) (rule is_inf_meet) + + lemma %invisible meet_le [intro?]: + "\ z \ x; z \ y \ \ z \ x \ y" + by (rule is_inf_greatest) (rule is_inf_meet) + + lemma %invisible join_equality [elim?]: "is_sup x y s \ x \ y = s" + proof (unfold join_def) + assume "is_sup x y s" + then show "(THE s. is_sup x y s) = s" + by (rule the_equality) (rule is_sup_uniq [OF _ `is_sup x y s`]) + qed + + lemma %invisible joinI [intro?]: "x \ s \ y \ s \ + (\z. x \ z \ y \ z \ s \ z) \ x \ y = s" + by (rule join_equality, rule is_supI) blast+ + + lemma %invisible is_sup_join [intro?]: "is_sup x y (x \ y)" + proof (unfold join_def) + from ex_sup obtain s where "is_sup x y s" .. + then show "is_sup x y (THE s. is_sup x y s)" + by (rule theI) (rule is_sup_uniq [OF _ `is_sup x y s`]) + qed + + lemma %invisible join_left [intro?]: + "x \ x \ y" + by (rule is_sup_upper) (rule is_sup_join) + + lemma %invisible join_right [intro?]: + "y \ x \ y" + by (rule is_sup_upper) (rule is_sup_join) + + lemma %invisible join_le [intro?]: + "\ x \ z; y \ z \ \ x \ y \ z" + by (rule is_sup_least) (rule is_sup_join) + + theorem %invisible meet_assoc: "(x \ y) \ z = x \ (y \ z)" + proof (rule meetI) + show "x \ (y \ z) \ x \ y" + proof + show "x \ (y \ z) \ x" .. + show "x \ (y \ z) \ y" + proof - + have "x \ (y \ z) \ y \ z" .. + also have "\ \ y" .. + finally show ?thesis . + qed + qed + show "x \ (y \ z) \ z" + proof - + have "x \ (y \ z) \ y \ z" .. + also have "\ \ z" .. + finally show ?thesis . + qed + fix w assume "w \ x \ y" and "w \ z" + show "w \ x \ (y \ z)" + proof + show "w \ x" + proof - + have "w \ x \ y" by fact + also have "\ \ x" .. + finally show ?thesis . + qed + show "w \ y \ z" + proof + show "w \ y" + proof - + have "w \ x \ y" by fact + also have "\ \ y" .. + finally show ?thesis . + qed + show "w \ z" by fact + qed + qed + qed + + theorem %invisible meet_commute: "x \ y = y \ x" + proof (rule meetI) + show "y \ x \ x" .. + show "y \ x \ y" .. + fix z assume "z \ y" and "z \ x" + then show "z \ y \ x" .. + qed + + theorem %invisible meet_join_absorb: "x \ (x \ y) = x" + proof (rule meetI) + show "x \ x" .. + show "x \ x \ y" .. + fix z assume "z \ x" and "z \ x \ y" + show "z \ x" by fact + qed + + theorem %invisible join_assoc: "(x \ y) \ z = x \ (y \ z)" + proof (rule joinI) + show "x \ y \ x \ (y \ z)" + proof + show "x \ x \ (y \ z)" .. + show "y \ x \ (y \ z)" + proof - + have "y \ y \ z" .. + also have "... \ x \ (y \ z)" .. + finally show ?thesis . + qed + qed + show "z \ x \ (y \ z)" + proof - + have "z \ y \ z" .. + also have "... \ x \ (y \ z)" .. + finally show ?thesis . + qed + fix w assume "x \ y \ w" and "z \ w" + show "x \ (y \ z) \ w" + proof + show "x \ w" + proof - + have "x \ x \ y" .. + also have "\ \ w" by fact + finally show ?thesis . + qed + show "y \ z \ w" + proof + show "y \ w" + proof - + have "y \ x \ y" .. + also have "... \ w" by fact + finally show ?thesis . + qed + show "z \ w" by fact + qed + qed + qed + + theorem %invisible join_commute: "x \ y = y \ x" + proof (rule joinI) + show "x \ y \ x" .. + show "y \ y \ x" .. + fix z assume "y \ z" and "x \ z" + then show "y \ x \ z" .. + qed + + theorem %invisible join_meet_absorb: "x \ (x \ y) = x" + proof (rule joinI) + show "x \ x" .. + show "x \ y \ x" .. + fix z assume "x \ z" and "x \ y \ z" + show "x \ z" by fact + qed + + theorem %invisible meet_idem: "x \ x = x" + proof - + have "x \ (x \ (x \ x)) = x" by (rule meet_join_absorb) + also have "x \ (x \ x) = x" by (rule join_meet_absorb) + finally show ?thesis . + qed + + theorem %invisible meet_related [elim?]: "x \ y \ x \ y = x" + proof (rule meetI) + assume "x \ y" + show "x \ x" .. + show "x \ y" by fact + fix z assume "z \ x" and "z \ y" + show "z \ x" by fact + qed + + theorem %invisible meet_related2 [elim?]: "y \ x \ x \ y = y" + by (drule meet_related) (simp add: meet_commute) + + theorem %invisible join_related [elim?]: "x \ y \ x \ y = y" + proof (rule joinI) + assume "x \ y" + show "y \ y" .. + show "x \ y" by fact + fix z assume "x \ z" and "y \ z" + show "y \ z" by fact + qed + + theorem %invisible join_related2 [elim?]: "y \ x \ x \ y = x" + by (drule join_related) (simp add: join_commute) + + theorem %invisible meet_connection: "(x \ y) = (x \ y = x)" + proof + assume "x \ y" + then have "is_inf x y x" .. + then show "x \ y = x" .. + next + have "x \ y \ y" .. + also assume "x \ y = x" + finally show "x \ y" . + qed + + theorem %invisible join_connection: "(x \ y) = (x \ y = y)" + proof + assume "x \ y" + then have "is_sup x y y" .. + then show "x \ y = y" .. + next + have "x \ x \ y" .. + also assume "x \ y = y" + finally show "x \ y" . + qed + + theorem %invisible meet_connection2: "(x \ y) = (y \ x = x)" + using meet_commute meet_connection by simp + + theorem %invisible join_connection2: "(x \ y) = (x \ y = y)" + using join_commute join_connection by simp + + text %invisible {* Naming according to Jacobson I, p.\ 459. *} + lemmas %invisible L1 = join_commute meet_commute + lemmas %invisible L2 = join_assoc meet_assoc + (* lemmas L3 = join_idem meet_idem *) + lemmas %invisible L4 = join_meet_absorb meet_join_absorb + + end + +text {* Locales for total orders and distributive lattices follow to + establish a sufficiently rich landscape of locales for + further examples in this tutorial. Each comes with an example + theorem. *} + + locale total_order = partial_order + + assumes total: "x \ y \ y \ x" + + lemma (in total_order) less_total: "x \ y \ x = y \ y \ x" + using total + by (unfold less_def) blast + + locale distrib_lattice = lattice + + assumes meet_distr: "x \ (y \ z) = x \ y \ x \ z" + + lemma (in distrib_lattice) join_distr: + "x \ (y \ z) = (x \ y) \ (x \ z)" (* txt {* Jacobson I, p.\ 462 *} *) + proof - + have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by (simp add: L4) + also have "... = x \ ((x \ z) \ (y \ z))" by (simp add: L2) + also have "... = x \ ((x \ y) \ z)" by (simp add: L1 meet_distr) + also have "... = ((x \ y) \ x) \ ((x \ y) \ z)" by (simp add: L1 L4) + also have "... = (x \ y) \ (x \ z)" by (simp add: meet_distr) + finally show ?thesis . + qed + +text {* + The locale hierarchy obtained through these declarations is shown in + Figure~\ref{fig:lattices}(a). + +\begin{figure} +\hrule \vspace{2ex} +\begin{center} +\subfigure[Declared hierarchy]{ +\begin{tikzpicture} + \node (po) at (0,0) {@{text partial_order}}; + \node (lat) at (-1.5,-1) {@{text lattice}}; + \node (dlat) at (-1.5,-2) {@{text distrib_lattice}}; + \node (to) at (1.5,-1) {@{text total_order}}; + \draw (po) -- (lat); + \draw (lat) -- (dlat); + \draw (po) -- (to); +% \draw[->, dashed] (lat) -- (to); +\end{tikzpicture} +} \\ +\subfigure[Total orders are lattices]{ +\begin{tikzpicture} + \node (po) at (0,0) {@{text partial_order}}; + \node (lat) at (0,-1) {@{text lattice}}; + \node (dlat) at (-1.5,-2) {@{text distrib_lattice}}; + \node (to) at (1.5,-2) {@{text total_order}}; + \draw (po) -- (lat); + \draw (lat) -- (dlat); + \draw (lat) -- (to); +% \draw[->, dashed] (dlat) -- (to); +\end{tikzpicture} +} \quad +\subfigure[Total orders are distributive lattices]{ +\begin{tikzpicture} + \node (po) at (0,0) {@{text partial_order}}; + \node (lat) at (0,-1) {@{text lattice}}; + \node (dlat) at (0,-2) {@{text distrib_lattice}}; + \node (to) at (0,-3) {@{text total_order}}; + \draw (po) -- (lat); + \draw (lat) -- (dlat); + \draw (dlat) -- (to); +\end{tikzpicture} +} +\end{center} +\hrule +\caption{Hierarchy of Lattice Locales.} +\label{fig:lattices} +\end{figure} + *} + +section {* Changing the Locale Hierarchy + \label{sec:changing-the-hierarchy} *} + +text {* + Locales enable to prove theorems abstractly, relative to + sets of assumptions. These theorems can then be used in other + contexts where the assumptions themselves, or + instances of the assumptions, are theorems. This form of theorem + reuse is called \emph{interpretation}. Locales generalise + interpretation from theorems to conclusions, enabling the reuse of + definitions and other constructs that are not part of the + specifications of the locales. + + The first form of interpretation we will consider in this tutorial + is provided by the \isakeyword{sublocale} command. It enables to + modify the import hierarchy to reflect the \emph{logical} relation + between locales. + + Consider the locale hierarchy from Figure~\ref{fig:lattices}(a). + Total orders are lattices, although this is not reflected here, and + definitions, theorems and other conclusions + from @{term lattice} are not available in @{term total_order}. To + obtain the situation in Figure~\ref{fig:lattices}(b), it is + sufficient to add the conclusions of the latter locale to the former. + The \isakeyword{sublocale} command does exactly this. + The declaration \isakeyword{sublocale} $l_1 + \subseteq l_2$ causes locale $l_2$ to be \emph{interpreted} in the + context of $l_1$. This means that all conclusions of $l_2$ are made + available in $l_1$. + + Of course, the change of hierarchy must be supported by a theorem + that reflects, in our example, that total orders are indeed + lattices. Therefore the \isakeyword{sublocale} command generates a + goal, which must be discharged by the user. This is illustrated in + the following paragraphs. First the sublocale relation is stated. +*} + + sublocale %visible total_order \ lattice + +txt {* \normalsize + This enters the context of locale @{text total_order}, in + which the goal @{subgoals [display]} must be shown. + Now the + locale predicate needs to be unfolded --- for example, using its + definition or by introduction rules + provided by the locale package. For automation, the locale package + provides the methods @{text intro_locales} and @{text + unfold_locales}. They are aware of the + current context and dependencies between locales and automatically + discharge goals implied by these. While @{text unfold_locales} + always unfolds locale predicates to assumptions, @{text + intro_locales} only unfolds definitions along the locale + hierarchy, leaving a goal consisting of predicates defined by the + locale package. Occasionally the latter is of advantage since the goal + is smaller. + + For the current goal, we would like to get hold of + the assumptions of @{text lattice}, which need to be shown, hence + @{text unfold_locales} is appropriate. *} + + proof unfold_locales + +txt {* \normalsize + Since the fact that both lattices and total orders are partial + orders is already reflected in the locale hierarchy, the assumptions + of @{text partial_order} are discharged automatically, and only the + assumptions introduced in @{text lattice} remain as subgoals + @{subgoals [display]} + The proof for the first subgoal is obtained by constructing an + infimum, whose existence is implied by totality. *} + + fix x y + from total have "is_inf x y (if x \ y then x else y)" + by (auto simp: is_inf_def) + then show "\inf. is_inf x y inf" .. +txt {* \normalsize + The proof for the second subgoal is analogous and not + reproduced here. *} + next %invisible + fix x y + from total have "is_sup x y (if x \ y then y else x)" + by (auto simp: is_sup_def) + then show "\sup. is_sup x y sup" .. qed %visible + +text {* Similarly, we may establish that total orders are distributive + lattices with a second \isakeyword{sublocale} statement. *} + + sublocale total_order \ distrib_lattice + proof unfold_locales + fix %"proof" x y z + show "x \ (y \ z) = x \ y \ x \ z" (is "?l = ?r") + txt {* Jacobson I, p.\ 462 *} + proof - + { assume c: "y \ x" "z \ x" + from c have "?l = y \ z" + by (metis c join_connection2 join_related2 meet_related2 total) + also from c have "... = ?r" by (metis meet_related2) + finally have "?l = ?r" . } + moreover + { assume c: "x \ y \ x \ z" + from c have "?l = x" + by (metis join_connection2 join_related2 meet_connection total trans) + also from c have "... = ?r" + by (metis join_commute join_related2 meet_connection meet_related2 total) + finally have "?l = ?r" . } + moreover note total + ultimately show ?thesis by blast + qed + qed + +text {* The locale hierarchy is now as shown in + Figure~\ref{fig:lattices}(c). *} + +text {* + Locale interpretation is \emph{dynamic}. The statement + \isakeyword{sublocale} $l_1 \subseteq l_2$ will not just add the + current conclusions of $l_2$ to $l_1$. Rather the dependency is + stored, and conclusions that will be + added to $l_2$ in future are automatically propagated to $l_1$. + The sublocale relation is transitive --- that is, propagation takes + effect along chains of sublocales. Even cycles in the sublocale relation are + supported, as long as these cycles do not lead to infinite chains. + Details are discussed in the technical report \cite{Ballarin2006a}. + See also Section~\ref{sec:infinite-chains} of this tutorial. *} + +end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Examples1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Examples1.thy Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,89 @@ +theory Examples1 +imports Examples +begin +text {* \vspace{-5ex} *} +section {* Use of Locales in Theories and Proofs + \label{sec:interpretation} *} + +text {* + Locales can be interpreted in the contexts of theories and + structured proofs. These interpretations are dynamic, too. + Conclusions of locales will be propagated to the current theory or + the current proof context.% +\footnote{Strictly speaking, only interpretation in theories is + dynamic since it is not possible to change locales or the locale + hierarchy from within a proof.} + The focus of this section is on + interpretation in theories, but we will also encounter + interpretations in proofs, in + Section~\ref{sec:local-interpretation}. + + As an example, consider the type of integers @{typ int}. The + relation @{term "op \"} is a total order over @{typ int}. We start + with the interpretation that @{term "op \"} is a partial order. The + facilities of the interpretation command are explored gradually in + three versions. + *} + + +subsection {* First Version: Replacement of Parameters Only + \label{sec:po-first} *} + +text {* + The command \isakeyword{interpretation} is for the interpretation of + locale in theories. In the following example, the parameter of locale + @{text partial_order} is replaced by @{term "op \ :: int \ int \ + bool"} and the locale instance is interpreted in the current + theory. *} + + interpretation %visible int: partial_order "op \ :: int \ int \ bool" +txt {* \normalsize + The argument of the command is a simple \emph{locale expression} + consisting of the name of the interpreted locale, which is + preceded by the qualifier @{text "int:"} and succeeded by a + white-space-separated list of terms, which provide a full + instantiation of the locale parameters. The parameters are referred + to by order of declaration, which is also the order in which + \isakeyword{print\_locale} outputs them. The locale has only a + single parameter, hence the list of instantiation terms is a + singleton. + + The command creates the goal + @{subgoals [display]} which can be shown easily: + *} + by unfold_locales auto + +text {* The effect of the command is that instances of all + conclusions of the locale are available in the theory, where names + are prefixed by the qualifier. For example, transitivity for @{typ + int} is named @{thm [source] int.trans} and is the following + theorem: + @{thm [display, indent=2] int.trans} + It is not possible to reference this theorem simply as @{text + trans}. This prevents unwanted hiding of existing theorems of the + theory by an interpretation. *} + + +subsection {* Second Version: Replacement of Definitions *} + +text {* Not only does the above interpretation qualify theorem names. + The prefix @{text int} is applied to all names introduced in locale + conclusions including names introduced in definitions. The + qualified name @{text int.less} is short for + the interpretation of the definition, which is @{term int.less}. + Qualified name and expanded form may be used almost + interchangeably.% +\footnote{Since @{term "op \"} is polymorphic, for @{term int.less} a + more general type will be inferred than for @{text int.less} which + is over type @{typ int}.} + The latter is preferred on output, as for example in the theorem + @{thm [source] int.less_le_trans}: @{thm [display, indent=2] + int.less_le_trans} + Both notations for the strict order are not satisfactory. The + constant @{term "op <"} is the strict order for @{typ int}. + In order to allow for the desired replacement, interpretation + accepts \emph{equations} in addition to the parameter instantiation. + These follow the locale expression and are indicated with the + keyword \isakeyword{where}. This is the revised interpretation: + *} +end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Examples2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Examples2.thy Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,24 @@ +theory Examples2 +imports Examples +begin +text {* \vspace{-5ex} *} + interpretation %visible int: partial_order "op \ :: [int, int] \ bool" + where "int.less x y = (x < y)" + proof - + txt {* \normalsize The goals are now: + @{subgoals [display]} + The proof that~@{text \} is a partial order is as above. *} + show "partial_order (op \ :: int \ int \ bool)" + by unfold_locales auto + txt {* \normalsize The second goal is shown by unfolding the + definition of @{term "partial_order.less"}. *} + show "partial_order.less op \ x y = (x < y)" + unfolding partial_order.less_def [OF `partial_order op \`] + by auto + qed + +text {* Note that the above proof is not in the context of the + interpreted locale. Hence, the premise of @{text + "partial_order.less_def"} is discharged manually with @{text OF}. + *} +end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Examples3.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Examples3.thy Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,659 @@ +theory Examples3 +imports Examples +begin +text {* \vspace{-5ex} *} +subsection {* Third Version: Local Interpretation + \label{sec:local-interpretation} *} + +text {* In the above example, the fact that @{term "op \"} is a partial + order for the integers was used in the second goal to + discharge the premise in the definition of @{text "op \"}. In + general, proofs of the equations not only may involve definitions + from the interpreted locale but arbitrarily complex arguments in the + context of the locale. Therefore it would be convenient to have the + interpreted locale conclusions temporarily available in the proof. + This can be achieved by a locale interpretation in the proof body. + The command for local interpretations is \isakeyword{interpret}. We + repeat the example from the previous section to illustrate this. *} + + interpretation %visible int: partial_order "op \ :: int \ int \ bool" + where "int.less x y = (x < y)" + proof - + show "partial_order (op \ :: int \ int \ bool)" + by unfold_locales auto + then interpret int: partial_order "op \ :: [int, int] \ bool" . + show "int.less x y = (x < y)" + unfolding int.less_def by auto + qed + +text {* The inner interpretation is immediate from the preceding fact + and proved by assumption (Isar short hand ``.''). It enriches the + local proof context by the theorems + also obtained in the interpretation from Section~\ref{sec:po-first}, + and @{text int.less_def} may directly be used to unfold the + definition. Theorems from the local interpretation disappear after + leaving the proof context --- that is, after the succeeding + \isakeyword{next} or \isakeyword{qed} statement. *} + + +subsection {* Further Interpretations *} + +text {* Further interpretations are necessary for + the other locales. In @{text lattice} the operations~@{text \} + and~@{text \} are substituted by @{term "min :: int \ int \ int"} + and @{term "max :: int \ int \ int"}. The entire proof for the + interpretation is reproduced to give an example of a more + elaborate interpretation proof. Note that the equations are named + so they can be used in a later example. *} + + interpretation %visible int: lattice "op \ :: int \ int \ bool" + where int_min_eq: "int.meet x y = min x y" + and int_max_eq: "int.join x y = max x y" + proof - + show "lattice (op \ :: int \ int \ bool)" + txt {* \normalsize We have already shown that this is a partial + order, *} + apply unfold_locales + txt {* \normalsize hence only the lattice axioms remain to be + shown. + @{subgoals [display]} + By @{text is_inf} and @{text is_sup}, *} + apply (unfold int.is_inf_def int.is_sup_def) + txt {* \normalsize the goals are transformed to these + statements: + @{subgoals [display]} + This is Presburger arithmetic, which can be solved by the + method @{text arith}. *} + by arith+ + txt {* \normalsize In order to show the equations, we put ourselves + in a situation where the lattice theorems can be used in a + convenient way. *} + then interpret int: lattice "op \ :: int \ int \ bool" . + show "int.meet x y = min x y" + by (bestsimp simp: int.meet_def int.is_inf_def) + show "int.join x y = max x y" + by (bestsimp simp: int.join_def int.is_sup_def) + qed + +text {* Next follows that @{text "op \"} is a total order, again for + the integers. *} + + interpretation %visible int: total_order "op \ :: int \ int \ bool" + by unfold_locales arith + +text {* Theorems that are available in the theory at this point are shown in + Table~\ref{tab:int-lattice}. Two points are worth noting: + +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{l} + @{thm [source] int.less_def} from locale @{text partial_order}: \\ + \quad @{thm int.less_def} \\ + @{thm [source] int.meet_left} from locale @{text lattice}: \\ + \quad @{thm int.meet_left} \\ + @{thm [source] int.join_distr} from locale @{text distrib_lattice}: \\ + \quad @{thm int.join_distr} \\ + @{thm [source] int.less_total} from locale @{text total_order}: \\ + \quad @{thm int.less_total} +\end{tabular} +\end{center} +\hrule +\caption{Interpreted theorems for~@{text \} on the integers.} +\label{tab:int-lattice} +\end{table} + +\begin{itemize} +\item + Locale @{text distrib_lattice} was also interpreted. Since the + locale hierarchy reflects that total orders are distributive + lattices, the interpretation of the latter was inserted + automatically with the interpretation of the former. In general, + interpretation traverses the locale hierarchy upwards and interprets + all encountered locales, regardless whether imported or proved via + the \isakeyword{sublocale} command. Existing interpretations are + skipped avoiding duplicate work. +\item + The predicate @{term "op <"} appears in theorem @{thm [source] + int.less_total} + although an equation for the replacement of @{text "op \"} was only + given in the interpretation of @{text partial_order}. The + interpretation equations are pushed downwards the hierarchy for + related interpretations --- that is, for interpretations that share + the instances of parameters they have in common. +\end{itemize} + *} + +text {* The interpretations for a locale $n$ within the current + theory may be inspected with \isakeyword{print\_interps}~$n$. This + prints the list of instances of $n$, for which interpretations exist. + For example, \isakeyword{print\_interps} @{term partial_order} + outputs the following: +\begin{small} +\begin{alltt} + int! : partial_order "op \(\le\)" +\end{alltt} +\end{small} + Of course, there is only one interpretation. + The interpretation qualifier on the left is decorated with an + exclamation point. This means that it is mandatory. Qualifiers + can either be \emph{mandatory} or \emph{optional}, designated by + ``!'' or ``?'' respectively. Mandatory qualifiers must occur in a + name reference while optional ones need not. Mandatory qualifiers + prevent accidental hiding of names, while optional qualifiers can be + more convenient to use. For \isakeyword{interpretation}, the + default is ``!''. +*} + + +section {* Locale Expressions \label{sec:expressions} *} + +text {* + A map~@{term \} between partial orders~@{text \} and~@{text \} + is called order preserving if @{text "x \ y"} implies @{text "\ x \ + \ y"}. This situation is more complex than those encountered so + far: it involves two partial orders, and it is desirable to use the + existing locale for both. + + A locale for order preserving maps requires three parameters: @{text + le}~(\isakeyword{infixl}~@{text \}) and @{text + le'}~(\isakeyword{infixl}~@{text \}) for the orders and~@{text \} + for the map. + + In order to reuse the existing locale for partial orders, which has + the single parameter~@{text le}, it must be imported twice, once + mapping its parameter to~@{text le} from the new locale and once + to~@{text le'}. This can be achieved with a compound locale + expression. + + In general, a locale expression is a sequence of \emph{locale instances} + separated by~``$\textbf{+}$'' and followed by a \isakeyword{for} + clause. + An instance has the following format: +\begin{quote} + \textit{qualifier} \textbf{:} \textit{locale-name} + \textit{parameter-instantiation} +\end{quote} + We have already seen locale instances as arguments to + \isakeyword{interpretation} in Section~\ref{sec:interpretation}. + As before, the qualifier serves to disambiguate names from + different instances of the same locale. While in + \isakeyword{interpretation} qualifiers default to mandatory, in + import and in the \isakeyword{sublocale} command, they default to + optional. + + Since the parameters~@{text le} and~@{text le'} are to be partial + orders, our locale for order preserving maps will import the these + instances: +\begin{small} +\begin{alltt} + le: partial_order le + le': partial_order le' +\end{alltt} +\end{small} + For matter of convenience we choose to name parameter names and + qualifiers alike. This is an arbitrary decision. Technically, qualifiers + and parameters are unrelated. + + Having determined the instances, let us turn to the \isakeyword{for} + clause. It serves to declare locale parameters in the same way as + the context element \isakeyword{fixes} does. Context elements can + only occur after the import section, and therefore the parameters + referred to in the instances must be declared in the \isakeyword{for} + clause. The \isakeyword{for} clause is also where the syntax of these + parameters is declared. + + Two context elements for the map parameter~@{text \} and the + assumptions that it is order preserving complete the locale + declaration. *} + + locale order_preserving = + le: partial_order le + le': partial_order le' + for le (infixl "\" 50) and le' (infixl "\" 50) + + fixes \ + assumes hom_le: "x \ y \ \ x \ \ y" + +text (in order_preserving) {* Here are examples of theorems that are + available in the locale: + + \hspace*{1em}@{thm [source] hom_le}: @{thm hom_le} + + \hspace*{1em}@{thm [source] le.less_le_trans}: @{thm le.less_le_trans} + + \hspace*{1em}@{thm [source] le'.less_le_trans}: + @{thm [display, indent=4] le'.less_le_trans} + While there is infix syntax for the strict operation associated to + @{term "op \"}, there is none for the strict version of @{term + "op \"}. The abbreviation @{text less} with its infix syntax is only + available for the original instance it was declared for. We may + introduce the abbreviation @{text less'} with infix syntax~@{text \} + with the following declaration: *} + + abbreviation (in order_preserving) + less' (infixl "\" 50) where "less' \ partial_order.less le'" + +text (in order_preserving) {* Now the theorem is displayed nicely as + @{thm [source] le'.less_le_trans}: + @{thm [display, indent=2] le'.less_le_trans} *} + +text {* There are short notations for locale expressions. These are + discussed in the following. *} + + +subsection {* Default Instantiations *} + +text {* + It is possible to omit parameter instantiations. The + instantiation then defaults to the name of + the parameter itself. For example, the locale expression @{text + partial_order} is short for @{text "partial_order le"}, since the + locale's single parameter is~@{text le}. We took advantage of this + in the \isakeyword{sublocale} declarations of + Section~\ref{sec:changing-the-hierarchy}. *} + + +subsection {* Implicit Parameters \label{sec:implicit-parameters} *} + +text {* In a locale expression that occurs within a locale + declaration, omitted parameters additionally extend the (possibly + empty) \isakeyword{for} clause. + + The \isakeyword{for} clause is a general construct of Isabelle/Isar + to mark names occurring in the preceding declaration as ``arbitrary + but fixed''. This is necessary for example, if the name is already + bound in a surrounding context. In a locale expression, names + occurring in parameter instantiations should be bound by a + \isakeyword{for} clause whenever these names are not introduced + elsewhere in the context --- for example, on the left hand side of a + \isakeyword{sublocale} declaration. + + There is an exception to this rule in locale declarations, where the + \isakeyword{for} clause serves to declare locale parameters. Here, + locale parameters for which no parameter instantiation is given are + implicitly added, with their mixfix syntax, at the beginning of the + \isakeyword{for} clause. For example, in a locale declaration, the + expression @{text partial_order} is short for +\begin{small} +\begin{alltt} + partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.} +\end{alltt} +\end{small} + This short hand was used in the locale declarations throughout + Section~\ref{sec:import}. + *} + +text{* + The following locale declarations provide more examples. A + map~@{text \} is a lattice homomorphism if it preserves meet and + join. *} + + locale lattice_hom = + le: lattice + le': lattice le' for le' (infixl "\" 50) + + fixes \ + assumes hom_meet: "\ (x \ y) = le'.meet (\ x) (\ y)" + and hom_join: "\ (x \ y) = le'.join (\ x) (\ y)" + +text {* The parameter instantiation in the first instance of @{term + lattice} is omitted. This causes the parameter~@{text le} to be + added to the \isakeyword{for} clause, and the locale has + parameters~@{text le},~@{text le'} and, of course,~@{text \}. + + Before turning to the second example, we complete the locale by + providing infix syntax for the meet and join operations of the + second lattice. +*} + + context lattice_hom + begin + abbreviation meet' (infixl "\''" 50) where "meet' \ le'.meet" + abbreviation join' (infixl "\''" 50) where "join' \ le'.join" + end + +text {* The next example makes radical use of the short hand + facilities. A homomorphism is an endomorphism if both orders + coincide. *} + + locale lattice_end = lattice_hom _ le + +text {* The notation~@{text _} enables to omit a parameter in a + positional instantiation. The omitted parameter,~@{text le} becomes + the parameter of the declared locale and is, in the following + position, used to instantiate the second parameter of @{text + lattice_hom}. The effect is that of identifying the first in second + parameter of the homomorphism locale. *} + +text {* The inheritance diagram of the situation we have now is shown + in Figure~\ref{fig:hom}, where the dashed line depicts an + interpretation which is introduced below. Parameter instantiations + are indicated by $\sqsubseteq \mapsto \preceq$ etc. By looking at + the inheritance diagram it would seem + that two identical copies of each of the locales @{text + partial_order} and @{text lattice} are imported by @{text + lattice_end}. This is not the case! Inheritance paths with + identical morphisms are automatically detected and + the conclusions of the respective locales appear only once. + +\begin{figure} +\hrule \vspace{2ex} +\begin{center} +\begin{tikzpicture} + \node (o) at (0,0) {@{text partial_order}}; + \node (oh) at (1.5,-2) {@{text order_preserving}}; + \node (oh1) at (1.5,-0.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; + \node (oh2) at (0,-1.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; + \node (l) at (-1.5,-2) {@{text lattice}}; + \node (lh) at (0,-4) {@{text lattice_hom}}; + \node (lh1) at (0,-2.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; + \node (lh2) at (-1.5,-3.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; + \node (le) at (0,-6) {@{text lattice_end}}; + \node (le1) at (0,-4.8) + [anchor=west]{$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; + \node (le2) at (0,-5.2) + [anchor=west]{$\scriptscriptstyle \preceq \mapsto \sqsubseteq$}; + \draw (o) -- (l); + \draw[dashed] (oh) -- (lh); + \draw (lh) -- (le); + \draw (o) .. controls (oh1.south west) .. (oh); + \draw (o) .. controls (oh2.north east) .. (oh); + \draw (l) .. controls (lh1.south west) .. (lh); + \draw (l) .. controls (lh2.north east) .. (lh); +\end{tikzpicture} +\end{center} +\hrule +\caption{Hierarchy of Homomorphism Locales.} +\label{fig:hom} +\end{figure} + *} + +text {* It can be shown easily that a lattice homomorphism is order + preserving. As the final example of this section, a locale + interpretation is used to assert this: *} + + sublocale lattice_hom \ order_preserving + proof unfold_locales + fix x y + assume "x \ y" + then have "y = (x \ y)" by (simp add: le.join_connection) + then have "\ y = (\ x \' \ y)" by (simp add: hom_join [symmetric]) + then show "\ x \ \ y" by (simp add: le'.join_connection) + qed + +text (in lattice_hom) {* + Theorems and other declarations --- syntax, in particular --- from + the locale @{text order_preserving} are now active in @{text + lattice_hom}, for example + @{thm [source] hom_le}: + @{thm [display, indent=2] hom_le} + This theorem will be useful in the following section. + *} + + +section {* Conditional Interpretation *} + +text {* There are situations where an interpretation is not possible + in the general case since the desired property is only valid if + certain conditions are fulfilled. Take, for example, the function + @{text "\i. n * i"} that scales its argument by a constant factor. + This function is order preserving (and even a lattice endomorphism) + with respect to @{term "op \"} provided @{text "n \ 0"}. + + It is not possible to express this using a global interpretation, + because it is in general unspecified whether~@{term n} is + non-negative, but one may make an interpretation in an inner context + of a proof where full information is available. + This is not fully satisfactory either, since potentially + interpretations may be required to make interpretations in many + contexts. What is + required is an interpretation that depends on the condition --- and + this can be done with the \isakeyword{sublocale} command. For this + purpose, we introduce a locale for the condition. *} + + locale non_negative = + fixes n :: int + assumes non_neg: "0 \ n" + +text {* It is again convenient to make the interpretation in an + incremental fashion, first for order preserving maps, the for + lattice endomorphisms. *} + + sublocale non_negative \ + order_preserving "op \" "op \" "\i. n * i" + using non_neg by unfold_locales (rule mult_left_mono) + +text {* While the proof of the previous interpretation + is straightforward from monotonicity lemmas for~@{term "op *"}, the + second proof follows a useful pattern. *} + + sublocale %visible non_negative \ lattice_end "op \" "\i. n * i" + proof (unfold_locales, unfold int_min_eq int_max_eq) + txt {* \normalsize Unfolding the locale predicates \emph{and} the + interpretation equations immediately yields two subgoals that + reflect the core conjecture. + @{subgoals [display]} + It is now necessary to show, in the context of @{term + non_negative}, that multiplication by~@{term n} commutes with + @{term min} and @{term max}. *} + qed (auto simp: hom_le) + +text (in order_preserving) {* The lemma @{thm [source] hom_le} + simplifies a proof that would have otherwise been lengthy and we may + consider making it a default rule for the simplifier: *} + + lemmas (in order_preserving) hom_le [simp] + + +subsection {* Avoiding Infinite Chains of Interpretations + \label{sec:infinite-chains} *} + +text {* Similar situations arise frequently in formalisations of + abstract algebra where it is desirable to express that certain + constructions preserve certain properties. For example, polynomials + over rings are rings, or --- an example from the domain where the + illustrations of this tutorial are taken from --- a partial order + may be obtained for a function space by point-wise lifting of the + partial order of the co-domain. This corresponds to the following + interpretation: *} + + sublocale %visible partial_order \ f: partial_order "\f g. \x. f x \ g x" + oops + +text {* Unfortunately this is a cyclic interpretation that leads to an + infinite chain, namely + @{text [display, indent=2] "partial_order \ partial_order (\f g. \x. f x \ g x) \ + partial_order (\f g. \x y. f x y \ g x y) \ \"} + and the interpretation is rejected. + + Instead it is necessary to declare a locale that is logically + equivalent to @{term partial_order} but serves to collect facts + about functions spaces where the co-domain is a partial order, and + to make the interpretation in its context: *} + + locale fun_partial_order = partial_order + + sublocale fun_partial_order \ + f: partial_order "\f g. \x. f x \ g x" + by unfold_locales (fast,rule,fast,blast intro: trans) + +text {* It is quite common in abstract algebra that such a construction + maps a hierarchy of algebraic structures (or specifications) to a + related hierarchy. By means of the same lifting, a function space + is a lattice if its co-domain is a lattice: *} + + locale fun_lattice = fun_partial_order + lattice + + sublocale fun_lattice \ f: lattice "\f g. \x. f x \ g x" + proof unfold_locales + fix f g + have "partial_order.is_inf (\f g. \x. f x \ g x) f g (\x. f x \ g x)" + apply (rule is_infI) apply rule+ apply (drule spec, assumption)+ done + then show "\inf. partial_order.is_inf (\f g. \x. f x \ g x) f g inf" + by fast + next + fix f g + have "partial_order.is_sup (\f g. \x. f x \ g x) f g (\x. f x \ g x)" + apply (rule is_supI) apply rule+ apply (drule spec, assumption)+ done + then show "\sup. partial_order.is_sup (\f g. \x. f x \ g x) f g sup" + by fast + qed + + +section {* Further Reading *} + +text {* More information on locales and their interpretation is + available. For the locale hierarchy of import and interpretation + dependencies see~\cite{Ballarin2006a}; interpretations in theories + and proofs are covered in~\cite{Ballarin2006b}. In the latter, I + show how interpretation in proofs enables to reason about families + of algebraic structures, which cannot be expressed with locales + directly. + + Haftmann and Wenzel~\cite{HaftmannWenzel2007} overcome a restriction + of axiomatic type classes through a combination with locale + interpretation. The result is a Haskell-style class system with a + facility to generate ML and Haskell code. Classes are sufficient for + simple specifications with a single type parameter. The locales for + orders and lattices presented in this tutorial fall into this + category. Order preserving maps, homomorphisms and vector spaces, + on the other hand, do not. + + The locales reimplementation for Isabelle 2009 provides, among other + improvements, a clean integration with Isabelle/Isar's local theory + mechanisms, which are described in another paper by Haftmann and + Wenzel~\cite{HaftmannWenzel2009}. + + The original work of Kamm\"uller on locales~\cite{KammullerEtAl1999} + may be of interest from a historical perspective. My previous + report on locales and locale expressions~\cite{Ballarin2004a} + describes a simpler form of expressions than available now and is + outdated. The mathematical background on orders and lattices is + taken from Jacobson's textbook on algebra~\cite[Chapter~8]{Jacobson1985}. + + The sources of this tutorial, which include all proofs, are + available with the Isabelle distribution at + \url{http://isabelle.in.tum.de}. + *} + +text {* +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{l>$c<$l} + \multicolumn{3}{l}{Miscellaneous} \\ + + \textit{attr-name} & ::= + & \textit{name} $|$ \textit{attribute} $|$ + \textit{name} \textit{attribute} \\ + \textit{qualifier} & ::= + & \textit{name} [``\textbf{?}'' $|$ ``\textbf{!}''] \\[2ex] + + \multicolumn{3}{l}{Context Elements} \\ + + \textit{fixes} & ::= + & \textit{name} [ ``\textbf{::}'' \textit{type} ] + [ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ + \textit{mixfix} ] \\ +\begin{comment} + \textit{constrains} & ::= + & \textit{name} ``\textbf{::}'' \textit{type} \\ +\end{comment} + \textit{assumes} & ::= + & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ +\begin{comment} + \textit{defines} & ::= + & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ + \textit{notes} & ::= + & [ \textit{attr-name} ``\textbf{=}'' ] + ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ +\end{comment} + + \textit{element} & ::= + & \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ +\begin{comment} + & | + & \textbf{constrains} \textit{constrains} + ( \textbf{and} \textit{constrains} )$^*$ \\ +\end{comment} + & | + & \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\[2ex] +%\begin{comment} +% & | +% & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ +% & | +% & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ +%\end{comment} + + \multicolumn{3}{l}{Locale Expressions} \\ + + \textit{pos-insts} & ::= + & ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\ + \textit{named-insts} & ::= + & \textbf{where} \textit{name} ``\textbf{=}'' \textit{term} + ( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\ + \textit{instance} & ::= + & [ \textit{qualifier} ``\textbf{:}'' ] + \textit{name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\ + \textit{expression} & ::= + & \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$ + [ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex] + + \multicolumn{3}{l}{Declaration of Locales} \\ + + \textit{locale} & ::= + & \textit{element}$^+$ \\ + & | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ + \textit{toplevel} & ::= + & \textbf{locale} \textit{name} [ ``\textbf{=}'' + \textit{locale} ] \\[2ex] + + \multicolumn{3}{l}{Interpretation} \\ + + \textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] + \textit{prop} \\ + \textit{equations} & ::= & \textbf{where} \textit{equation} ( \textbf{and} + \textit{equation} )$^*$ \\ + \textit{toplevel} & ::= + & \textbf{sublocale} \textit{name} ( ``$<$'' $|$ + ``$\subseteq$'' ) \textit{expression} \textit{proof} \\ + & | + & \textbf{interpretation} + \textit{expression} [ \textit{equations} ] \textit{proof} \\ + & | + & \textbf{interpret} + \textit{expression} \textit{proof} \\[2ex] + + \multicolumn{3}{l}{Diagnostics} \\ + + \textit{toplevel} & ::= + & \textbf{print\_locales} \\ + & | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{name} \\ + & | & \textbf{print\_interps} \textit{name} +\end{tabular} +\end{center} +\hrule +\caption{Syntax of Locale Commands.} +\label{tab:commands} +\end{table} + *} + +text {* \textbf{Revision History.} For the present third revision of + the tutorial, much of the explanatory text + was rewritten. Inheritance of interpretation equations is + available with the forthcoming release of Isabelle, which at the + time of editing these notes is expected for the end of 2009. + The second revision accommodates changes introduced by the locales + reimplementation for Isabelle 2009. Most notably locale expressions + have been generalised from renaming to instantiation. *} + +text {* \textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, + Randy Pollack, Andreas Schropp, Christian Sternagel and Makarius Wenzel + have made + useful comments on earlier versions of this document. The section + on conditional interpretation was inspired by a number of e-mail + enquiries the author received from locale users, and which suggested + that this use case is important enough to deserve explicit + explanation. The term \emph{conditional interpretation} is due to + Larry Paulson. *} + +end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/Examples.thy --- a/doc-src/Locales/Locales/Examples.thy Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,770 +0,0 @@ -theory Examples -imports Main -begin - -pretty_setmargin %invisible 65 - -(* -text {* The following presentation will use notation of - Isabelle's meta logic, hence a few sentences to explain this. - The logical - primitives are universal quantification (@{text "\"}), entailment - (@{text "\"}) and equality (@{text "\"}). Variables (not bound - variables) are sometimes preceded by a question mark. The logic is - typed. Type variables are denoted by~@{text "'a"},~@{text "'b"} - etc., and~@{text "\"} is the function type. Double brackets~@{text - "\"} and~@{text "\"} are used to abbreviate nested entailment. -*} -*) - -section {* Introduction *} - -text {* - Locales are based on contexts. A \emph{context} can be seen as a - formula schema -\[ - @{text "\x\<^sub>1\x\<^sub>n. \ A\<^sub>1; \ ;A\<^sub>m \ \ \"} -\] - where the variables~@{text "x\<^sub>1"}, \ldots,~@{text "x\<^sub>n"} are called - \emph{parameters} and the premises $@{text "A\<^sub>1"}, \ldots,~@{text - "A\<^sub>m"}$ \emph{assumptions}. A formula~@{text "C"} - is a \emph{theorem} in the context if it is a conclusion -\[ - @{text "\x\<^sub>1\x\<^sub>n. \ A\<^sub>1; \ ;A\<^sub>m \ \ C"}. -\] - Isabelle/Isar's notion of context goes beyond this logical view. - Its contexts record, in a consecutive order, proved - conclusions along with \emph{attributes}, which can provide context - specific configuration information for proof procedures and concrete - syntax. From a logical perspective, locales are just contexts that - have been made persistent. To the user, though, they provide - powerful means for declaring and combining contexts, and for the - reuse of theorems proved in these contexts. - *} - -section {* Simple Locales *} - -text {* - In its simplest form, a - \emph{locale declaration} consists of a sequence of context elements - declaring parameters (keyword \isakeyword{fixes}) and assumptions - (keyword \isakeyword{assumes}). The following is the specification of - partial orders, as locale @{text partial_order}. - *} - - locale partial_order = - fixes le :: "'a \ 'a \ bool" (infixl "\" 50) - assumes refl [intro, simp]: "x \ x" - and anti_sym [intro]: "\ x \ y; y \ x \ \ x = y" - and trans [trans]: "\ x \ y; y \ z \ \ x \ z" - -text (in partial_order) {* The parameter of this locale is~@{text le}, - which is a binary predicate with infix syntax~@{text \}. The - parameter syntax is available in the subsequent - assumptions, which are the familiar partial order axioms. - - Isabelle recognises unbound names as free variables. In locale - assumptions, these are implicitly universally quantified. That is, - @{term "\ x \ y; y \ z \ \ x \ z"} in fact means - @{term "\x y z. \ x \ y; y \ z \ \ x \ z"}. - - Two commands are provided to inspect locales: - \isakeyword{print\_locales} lists the names of all locales of the - current theory; \isakeyword{print\_locale}~$n$ prints the parameters - and assumptions of locale $n$; the variation \isakeyword{print\_locale!}~$n$ - additionally outputs the conclusions that are stored in the locale. - We may inspect the new locale - by issuing \isakeyword{print\_locale!} @{term partial_order}. The output - is the following list of context elements. -\begin{small} -\begin{alltt} - \isakeyword{fixes} le :: "'a \(\Rightarrow\) 'a \(\Rightarrow\) bool" (\isakeyword{infixl} "\(\sqsubseteq\)" 50) - \isakeyword{assumes} "partial_order op \(\sqsubseteq\)" - \isakeyword{notes} assumption - refl [intro, simp] = `?x \(\sqsubseteq\) ?x` - \isakeyword{and} - anti_sym [intro] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?x\(\isasymrbrakk\) \(\Longrightarrow\) ?x = ?y` - \isakeyword{and} - trans [trans] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?z\(\isasymrbrakk\) \(\Longrightarrow\) ?x \(\sqsubseteq\) ?z` -\end{alltt} -\end{small} - The keyword \isakeyword{notes} denotes a conclusion element. There - is one conclusion, which was added automatically. Instead, there is - only one assumption, namely @{term "partial_order le"}. The locale - declaration has introduced the predicate @{term - partial_order} to the theory. This predicate is the - \emph{locale predicate}. Its definition may be inspected by - issuing \isakeyword{thm} @{thm [source] partial_order_def}. - @{thm [display, indent=2] partial_order_def} - In our example, this is a unary predicate over the parameter of the - locale. It is equivalent to the original assumptions, which have - been turned into conclusions and are - available as theorems in the context of the locale. The names and - attributes from the locale declaration are associated to these - theorems and are effective in the context of the locale. - - Each conclusion has a \emph{foundational theorem} as counterpart - in the theory. Technically, this is simply the theorem composed - of context and conclusion. For the transitivity theorem, this is - @{thm [source] partial_order.trans}: - @{thm [display, indent=2] partial_order.trans} -*} - -subsection {* Targets: Extending Locales *} - -text {* - The specification of a locale is fixed, but its list of conclusions - may be extended through Isar commands that take a \emph{target} argument. - In the following, \isakeyword{definition} and - \isakeyword{theorem} are illustrated. - Table~\ref{tab:commands-with-target} lists Isar commands that accept - a target. Isar provides various ways of specifying the target. A - target for a single command may be indicated with keyword - \isakeyword{in} in the following way: - -\begin{table} -\hrule -\vspace{2ex} -\begin{center} -\begin{tabular}{ll} - \isakeyword{definition} & definition through an equation \\ - \isakeyword{inductive} & inductive definition \\ - \isakeyword{primrec} & primitive recursion \\ - \isakeyword{fun}, \isakeyword{function} & general recursion \\ - \isakeyword{abbreviation} & syntactic abbreviation \\ - \isakeyword{theorem}, etc.\ & theorem statement with proof \\ - \isakeyword{theorems}, etc.\ & redeclaration of theorems \\ - \isakeyword{text}, etc.\ & document markup -\end{tabular} -\end{center} -\hrule -\caption{Isar commands that accept a target.} -\label{tab:commands-with-target} -\end{table} - *} - - definition (in partial_order) - less :: "'a \ 'a \ bool" (infixl "\" 50) - where "(x \ y) = (x \ y \ x \ y)" - -text (in partial_order) {* The strict order @{text less} with infix - syntax~@{text \} is - defined in terms of the locale parameter~@{text le} and the general - equality of the object logic we work in. The definition generates a - \emph{foundational constant} - @{term partial_order.less} with definition @{thm [source] - partial_order.less_def}: - @{thm [display, indent=2] partial_order.less_def} - At the same time, the locale is extended by syntax transformations - hiding this construction in the context of the locale. Here, the - abbreviation @{text less} is available for - @{text "partial_order.less le"}, and it is printed - and parsed as infix~@{text \}. Finally, the conclusion @{thm [source] - less_def} is added to the locale: - @{thm [display, indent=2] less_def} -*} - -text {* The treatment of theorem statements is more straightforward. - As an example, here is the derivation of a transitivity law for the - strict order relation. *} - - lemma (in partial_order) less_le_trans [trans]: - "\ x \ y; y \ z \ \ x \ z" - unfolding %visible less_def by %visible (blast intro: trans) - -text {* In the context of the proof, conclusions of the - locale may be used like theorems. Attributes are effective: @{text - anti_sym} was - declared as introduction rule, hence it is in the context's set of - rules used by the classical reasoner by default. *} - -subsection {* Context Blocks *} - -text {* When working with locales, sequences of commands with the same - target are frequent. A block of commands, delimited by - \isakeyword{begin} and \isakeyword{end}, makes a theory-like style - of working possible. All commands inside the block refer to the - same target. A block may immediately follow a locale - declaration, which makes that locale the target. Alternatively the - target for a block may be given with the \isakeyword{context} - command. - - This style of working is illustrated in the block below, where - notions of infimum and supremum for partial orders are introduced, - together with theorems about their uniqueness. *} - - context partial_order - begin - - definition - is_inf where "is_inf x y i = - (i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i))" - - definition - is_sup where "is_sup x y s = - (x \ s \ y \ s \ (\z. x \ z \ y \ z \ s \ z))" - - lemma %invisible is_infI [intro?]: "i \ x \ i \ y \ - (\z. z \ x \ z \ y \ z \ i) \ is_inf x y i" - by (unfold is_inf_def) blast - - lemma %invisible is_inf_lower [elim?]: - "is_inf x y i \ (i \ x \ i \ y \ C) \ C" - by (unfold is_inf_def) blast - - lemma %invisible is_inf_greatest [elim?]: - "is_inf x y i \ z \ x \ z \ y \ z \ i" - by (unfold is_inf_def) blast - - theorem is_inf_uniq: "\is_inf x y i; is_inf x y i'\ \ i = i'" - proof - - assume inf: "is_inf x y i" - assume inf': "is_inf x y i'" - show ?thesis - proof (rule anti_sym) - from inf' show "i \ i'" - proof (rule is_inf_greatest) - from inf show "i \ x" .. - from inf show "i \ y" .. - qed - from inf show "i' \ i" - proof (rule is_inf_greatest) - from inf' show "i' \ x" .. - from inf' show "i' \ y" .. - qed - qed - qed - - theorem %invisible is_inf_related [elim?]: "x \ y \ is_inf x y x" - proof - - assume "x \ y" - show ?thesis - proof - show "x \ x" .. - show "x \ y" by fact - fix z assume "z \ x" and "z \ y" show "z \ x" by fact - qed - qed - - lemma %invisible is_supI [intro?]: "x \ s \ y \ s \ - (\z. x \ z \ y \ z \ s \ z) \ is_sup x y s" - by (unfold is_sup_def) blast - - lemma %invisible is_sup_least [elim?]: - "is_sup x y s \ x \ z \ y \ z \ s \ z" - by (unfold is_sup_def) blast - - lemma %invisible is_sup_upper [elim?]: - "is_sup x y s \ (x \ s \ y \ s \ C) \ C" - by (unfold is_sup_def) blast - - theorem is_sup_uniq: "\is_sup x y s; is_sup x y s'\ \ s = s'" - proof - - assume sup: "is_sup x y s" - assume sup': "is_sup x y s'" - show ?thesis - proof (rule anti_sym) - from sup show "s \ s'" - proof (rule is_sup_least) - from sup' show "x \ s'" .. - from sup' show "y \ s'" .. - qed - from sup' show "s' \ s" - proof (rule is_sup_least) - from sup show "x \ s" .. - from sup show "y \ s" .. - qed - qed - qed - - theorem %invisible is_sup_related [elim?]: "x \ y \ is_sup x y y" - proof - - assume "x \ y" - show ?thesis - proof - show "x \ y" by fact - show "y \ y" .. - fix z assume "x \ z" and "y \ z" - show "y \ z" by fact - qed - qed - - end - -text {* The syntax of the locale commands discussed in this tutorial is - shown in Table~\ref{tab:commands}. The grammar is complete with the - exception of the context elements \isakeyword{constrains} and - \isakeyword{defines}, which are provided for backward - compatibility. See the Isabelle/Isar Reference - Manual~\cite{IsarRef} for full documentation. *} - - -section {* Import \label{sec:import} *} - -text {* - Algebraic structures are commonly defined by adding operations and - properties to existing structures. For example, partial orders - are extended to lattices and total orders. Lattices are extended to - distributive lattices. *} - -text {* - With locales, this kind of inheritance is achieved through - \emph{import} of locales. The import part of a locale declaration, - if present, precedes the context elements. Here is an example, - where partial orders are extended to lattices. - *} - - locale lattice = partial_order + - assumes ex_inf: "\inf. is_inf x y inf" - and ex_sup: "\sup. is_sup x y sup" - begin - -text {* These assumptions refer to the predicates for infimum - and supremum defined for @{text partial_order} in the previous - section. We now introduce the notions of meet and join. *} - - definition - meet (infixl "\" 70) where "x \ y = (THE inf. is_inf x y inf)" - definition - join (infixl "\" 65) where "x \ y = (THE sup. is_sup x y sup)" - - lemma %invisible meet_equality [elim?]: "is_inf x y i \ x \ y = i" - proof (unfold meet_def) - assume "is_inf x y i" - then show "(THE i. is_inf x y i) = i" - by (rule the_equality) (rule is_inf_uniq [OF _ `is_inf x y i`]) - qed - - lemma %invisible meetI [intro?]: - "i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i) \ x \ y = i" - by (rule meet_equality, rule is_infI) blast+ - - lemma %invisible is_inf_meet [intro?]: "is_inf x y (x \ y)" - proof (unfold meet_def) - from ex_inf obtain i where "is_inf x y i" .. - then show "is_inf x y (THE i. is_inf x y i)" - by (rule theI) (rule is_inf_uniq [OF _ `is_inf x y i`]) - qed - - lemma %invisible meet_left [intro?]: - "x \ y \ x" - by (rule is_inf_lower) (rule is_inf_meet) - - lemma %invisible meet_right [intro?]: - "x \ y \ y" - by (rule is_inf_lower) (rule is_inf_meet) - - lemma %invisible meet_le [intro?]: - "\ z \ x; z \ y \ \ z \ x \ y" - by (rule is_inf_greatest) (rule is_inf_meet) - - lemma %invisible join_equality [elim?]: "is_sup x y s \ x \ y = s" - proof (unfold join_def) - assume "is_sup x y s" - then show "(THE s. is_sup x y s) = s" - by (rule the_equality) (rule is_sup_uniq [OF _ `is_sup x y s`]) - qed - - lemma %invisible joinI [intro?]: "x \ s \ y \ s \ - (\z. x \ z \ y \ z \ s \ z) \ x \ y = s" - by (rule join_equality, rule is_supI) blast+ - - lemma %invisible is_sup_join [intro?]: "is_sup x y (x \ y)" - proof (unfold join_def) - from ex_sup obtain s where "is_sup x y s" .. - then show "is_sup x y (THE s. is_sup x y s)" - by (rule theI) (rule is_sup_uniq [OF _ `is_sup x y s`]) - qed - - lemma %invisible join_left [intro?]: - "x \ x \ y" - by (rule is_sup_upper) (rule is_sup_join) - - lemma %invisible join_right [intro?]: - "y \ x \ y" - by (rule is_sup_upper) (rule is_sup_join) - - lemma %invisible join_le [intro?]: - "\ x \ z; y \ z \ \ x \ y \ z" - by (rule is_sup_least) (rule is_sup_join) - - theorem %invisible meet_assoc: "(x \ y) \ z = x \ (y \ z)" - proof (rule meetI) - show "x \ (y \ z) \ x \ y" - proof - show "x \ (y \ z) \ x" .. - show "x \ (y \ z) \ y" - proof - - have "x \ (y \ z) \ y \ z" .. - also have "\ \ y" .. - finally show ?thesis . - qed - qed - show "x \ (y \ z) \ z" - proof - - have "x \ (y \ z) \ y \ z" .. - also have "\ \ z" .. - finally show ?thesis . - qed - fix w assume "w \ x \ y" and "w \ z" - show "w \ x \ (y \ z)" - proof - show "w \ x" - proof - - have "w \ x \ y" by fact - also have "\ \ x" .. - finally show ?thesis . - qed - show "w \ y \ z" - proof - show "w \ y" - proof - - have "w \ x \ y" by fact - also have "\ \ y" .. - finally show ?thesis . - qed - show "w \ z" by fact - qed - qed - qed - - theorem %invisible meet_commute: "x \ y = y \ x" - proof (rule meetI) - show "y \ x \ x" .. - show "y \ x \ y" .. - fix z assume "z \ y" and "z \ x" - then show "z \ y \ x" .. - qed - - theorem %invisible meet_join_absorb: "x \ (x \ y) = x" - proof (rule meetI) - show "x \ x" .. - show "x \ x \ y" .. - fix z assume "z \ x" and "z \ x \ y" - show "z \ x" by fact - qed - - theorem %invisible join_assoc: "(x \ y) \ z = x \ (y \ z)" - proof (rule joinI) - show "x \ y \ x \ (y \ z)" - proof - show "x \ x \ (y \ z)" .. - show "y \ x \ (y \ z)" - proof - - have "y \ y \ z" .. - also have "... \ x \ (y \ z)" .. - finally show ?thesis . - qed - qed - show "z \ x \ (y \ z)" - proof - - have "z \ y \ z" .. - also have "... \ x \ (y \ z)" .. - finally show ?thesis . - qed - fix w assume "x \ y \ w" and "z \ w" - show "x \ (y \ z) \ w" - proof - show "x \ w" - proof - - have "x \ x \ y" .. - also have "\ \ w" by fact - finally show ?thesis . - qed - show "y \ z \ w" - proof - show "y \ w" - proof - - have "y \ x \ y" .. - also have "... \ w" by fact - finally show ?thesis . - qed - show "z \ w" by fact - qed - qed - qed - - theorem %invisible join_commute: "x \ y = y \ x" - proof (rule joinI) - show "x \ y \ x" .. - show "y \ y \ x" .. - fix z assume "y \ z" and "x \ z" - then show "y \ x \ z" .. - qed - - theorem %invisible join_meet_absorb: "x \ (x \ y) = x" - proof (rule joinI) - show "x \ x" .. - show "x \ y \ x" .. - fix z assume "x \ z" and "x \ y \ z" - show "x \ z" by fact - qed - - theorem %invisible meet_idem: "x \ x = x" - proof - - have "x \ (x \ (x \ x)) = x" by (rule meet_join_absorb) - also have "x \ (x \ x) = x" by (rule join_meet_absorb) - finally show ?thesis . - qed - - theorem %invisible meet_related [elim?]: "x \ y \ x \ y = x" - proof (rule meetI) - assume "x \ y" - show "x \ x" .. - show "x \ y" by fact - fix z assume "z \ x" and "z \ y" - show "z \ x" by fact - qed - - theorem %invisible meet_related2 [elim?]: "y \ x \ x \ y = y" - by (drule meet_related) (simp add: meet_commute) - - theorem %invisible join_related [elim?]: "x \ y \ x \ y = y" - proof (rule joinI) - assume "x \ y" - show "y \ y" .. - show "x \ y" by fact - fix z assume "x \ z" and "y \ z" - show "y \ z" by fact - qed - - theorem %invisible join_related2 [elim?]: "y \ x \ x \ y = x" - by (drule join_related) (simp add: join_commute) - - theorem %invisible meet_connection: "(x \ y) = (x \ y = x)" - proof - assume "x \ y" - then have "is_inf x y x" .. - then show "x \ y = x" .. - next - have "x \ y \ y" .. - also assume "x \ y = x" - finally show "x \ y" . - qed - - theorem %invisible join_connection: "(x \ y) = (x \ y = y)" - proof - assume "x \ y" - then have "is_sup x y y" .. - then show "x \ y = y" .. - next - have "x \ x \ y" .. - also assume "x \ y = y" - finally show "x \ y" . - qed - - theorem %invisible meet_connection2: "(x \ y) = (y \ x = x)" - using meet_commute meet_connection by simp - - theorem %invisible join_connection2: "(x \ y) = (x \ y = y)" - using join_commute join_connection by simp - - text %invisible {* Naming according to Jacobson I, p.\ 459. *} - lemmas %invisible L1 = join_commute meet_commute - lemmas %invisible L2 = join_assoc meet_assoc - (* lemmas L3 = join_idem meet_idem *) - lemmas %invisible L4 = join_meet_absorb meet_join_absorb - - end - -text {* Locales for total orders and distributive lattices follow to - establish a sufficiently rich landscape of locales for - further examples in this tutorial. Each comes with an example - theorem. *} - - locale total_order = partial_order + - assumes total: "x \ y \ y \ x" - - lemma (in total_order) less_total: "x \ y \ x = y \ y \ x" - using total - by (unfold less_def) blast - - locale distrib_lattice = lattice + - assumes meet_distr: "x \ (y \ z) = x \ y \ x \ z" - - lemma (in distrib_lattice) join_distr: - "x \ (y \ z) = (x \ y) \ (x \ z)" (* txt {* Jacobson I, p.\ 462 *} *) - proof - - have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by (simp add: L4) - also have "... = x \ ((x \ z) \ (y \ z))" by (simp add: L2) - also have "... = x \ ((x \ y) \ z)" by (simp add: L1 meet_distr) - also have "... = ((x \ y) \ x) \ ((x \ y) \ z)" by (simp add: L1 L4) - also have "... = (x \ y) \ (x \ z)" by (simp add: meet_distr) - finally show ?thesis . - qed - -text {* - The locale hierarchy obtained through these declarations is shown in - Figure~\ref{fig:lattices}(a). - -\begin{figure} -\hrule \vspace{2ex} -\begin{center} -\subfigure[Declared hierarchy]{ -\begin{tikzpicture} - \node (po) at (0,0) {@{text partial_order}}; - \node (lat) at (-1.5,-1) {@{text lattice}}; - \node (dlat) at (-1.5,-2) {@{text distrib_lattice}}; - \node (to) at (1.5,-1) {@{text total_order}}; - \draw (po) -- (lat); - \draw (lat) -- (dlat); - \draw (po) -- (to); -% \draw[->, dashed] (lat) -- (to); -\end{tikzpicture} -} \\ -\subfigure[Total orders are lattices]{ -\begin{tikzpicture} - \node (po) at (0,0) {@{text partial_order}}; - \node (lat) at (0,-1) {@{text lattice}}; - \node (dlat) at (-1.5,-2) {@{text distrib_lattice}}; - \node (to) at (1.5,-2) {@{text total_order}}; - \draw (po) -- (lat); - \draw (lat) -- (dlat); - \draw (lat) -- (to); -% \draw[->, dashed] (dlat) -- (to); -\end{tikzpicture} -} \quad -\subfigure[Total orders are distributive lattices]{ -\begin{tikzpicture} - \node (po) at (0,0) {@{text partial_order}}; - \node (lat) at (0,-1) {@{text lattice}}; - \node (dlat) at (0,-2) {@{text distrib_lattice}}; - \node (to) at (0,-3) {@{text total_order}}; - \draw (po) -- (lat); - \draw (lat) -- (dlat); - \draw (dlat) -- (to); -\end{tikzpicture} -} -\end{center} -\hrule -\caption{Hierarchy of Lattice Locales.} -\label{fig:lattices} -\end{figure} - *} - -section {* Changing the Locale Hierarchy - \label{sec:changing-the-hierarchy} *} - -text {* - Locales enable to prove theorems abstractly, relative to - sets of assumptions. These theorems can then be used in other - contexts where the assumptions themselves, or - instances of the assumptions, are theorems. This form of theorem - reuse is called \emph{interpretation}. Locales generalise - interpretation from theorems to conclusions, enabling the reuse of - definitions and other constructs that are not part of the - specifications of the locales. - - The first form of interpretation we will consider in this tutorial - is provided by the \isakeyword{sublocale} command. It enables to - modify the import hierarchy to reflect the \emph{logical} relation - between locales. - - Consider the locale hierarchy from Figure~\ref{fig:lattices}(a). - Total orders are lattices, although this is not reflected here, and - definitions, theorems and other conclusions - from @{term lattice} are not available in @{term total_order}. To - obtain the situation in Figure~\ref{fig:lattices}(b), it is - sufficient to add the conclusions of the latter locale to the former. - The \isakeyword{sublocale} command does exactly this. - The declaration \isakeyword{sublocale} $l_1 - \subseteq l_2$ causes locale $l_2$ to be \emph{interpreted} in the - context of $l_1$. This means that all conclusions of $l_2$ are made - available in $l_1$. - - Of course, the change of hierarchy must be supported by a theorem - that reflects, in our example, that total orders are indeed - lattices. Therefore the \isakeyword{sublocale} command generates a - goal, which must be discharged by the user. This is illustrated in - the following paragraphs. First the sublocale relation is stated. -*} - - sublocale %visible total_order \ lattice - -txt {* \normalsize - This enters the context of locale @{text total_order}, in - which the goal @{subgoals [display]} must be shown. - Now the - locale predicate needs to be unfolded --- for example, using its - definition or by introduction rules - provided by the locale package. For automation, the locale package - provides the methods @{text intro_locales} and @{text - unfold_locales}. They are aware of the - current context and dependencies between locales and automatically - discharge goals implied by these. While @{text unfold_locales} - always unfolds locale predicates to assumptions, @{text - intro_locales} only unfolds definitions along the locale - hierarchy, leaving a goal consisting of predicates defined by the - locale package. Occasionally the latter is of advantage since the goal - is smaller. - - For the current goal, we would like to get hold of - the assumptions of @{text lattice}, which need to be shown, hence - @{text unfold_locales} is appropriate. *} - - proof unfold_locales - -txt {* \normalsize - Since the fact that both lattices and total orders are partial - orders is already reflected in the locale hierarchy, the assumptions - of @{text partial_order} are discharged automatically, and only the - assumptions introduced in @{text lattice} remain as subgoals - @{subgoals [display]} - The proof for the first subgoal is obtained by constructing an - infimum, whose existence is implied by totality. *} - - fix x y - from total have "is_inf x y (if x \ y then x else y)" - by (auto simp: is_inf_def) - then show "\inf. is_inf x y inf" .. -txt {* \normalsize - The proof for the second subgoal is analogous and not - reproduced here. *} - next %invisible - fix x y - from total have "is_sup x y (if x \ y then y else x)" - by (auto simp: is_sup_def) - then show "\sup. is_sup x y sup" .. qed %visible - -text {* Similarly, we may establish that total orders are distributive - lattices with a second \isakeyword{sublocale} statement. *} - - sublocale total_order \ distrib_lattice - proof unfold_locales - fix %"proof" x y z - show "x \ (y \ z) = x \ y \ x \ z" (is "?l = ?r") - txt {* Jacobson I, p.\ 462 *} - proof - - { assume c: "y \ x" "z \ x" - from c have "?l = y \ z" - by (metis c join_connection2 join_related2 meet_related2 total) - also from c have "... = ?r" by (metis meet_related2) - finally have "?l = ?r" . } - moreover - { assume c: "x \ y \ x \ z" - from c have "?l = x" - by (metis join_connection2 join_related2 meet_connection total trans) - also from c have "... = ?r" - by (metis join_commute join_related2 meet_connection meet_related2 total) - finally have "?l = ?r" . } - moreover note total - ultimately show ?thesis by blast - qed - qed - -text {* The locale hierarchy is now as shown in - Figure~\ref{fig:lattices}(c). *} - -text {* - Locale interpretation is \emph{dynamic}. The statement - \isakeyword{sublocale} $l_1 \subseteq l_2$ will not just add the - current conclusions of $l_2$ to $l_1$. Rather the dependency is - stored, and conclusions that will be - added to $l_2$ in future are automatically propagated to $l_1$. - The sublocale relation is transitive --- that is, propagation takes - effect along chains of sublocales. Even cycles in the sublocale relation are - supported, as long as these cycles do not lead to infinite chains. - Details are discussed in the technical report \cite{Ballarin2006a}. - See also Section~\ref{sec:infinite-chains} of this tutorial. *} - -end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/Examples1.thy --- a/doc-src/Locales/Locales/Examples1.thy Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,89 +0,0 @@ -theory Examples1 -imports Examples -begin -text {* \vspace{-5ex} *} -section {* Use of Locales in Theories and Proofs - \label{sec:interpretation} *} - -text {* - Locales can be interpreted in the contexts of theories and - structured proofs. These interpretations are dynamic, too. - Conclusions of locales will be propagated to the current theory or - the current proof context.% -\footnote{Strictly speaking, only interpretation in theories is - dynamic since it is not possible to change locales or the locale - hierarchy from within a proof.} - The focus of this section is on - interpretation in theories, but we will also encounter - interpretations in proofs, in - Section~\ref{sec:local-interpretation}. - - As an example, consider the type of integers @{typ int}. The - relation @{term "op \"} is a total order over @{typ int}. We start - with the interpretation that @{term "op \"} is a partial order. The - facilities of the interpretation command are explored gradually in - three versions. - *} - - -subsection {* First Version: Replacement of Parameters Only - \label{sec:po-first} *} - -text {* - The command \isakeyword{interpretation} is for the interpretation of - locale in theories. In the following example, the parameter of locale - @{text partial_order} is replaced by @{term "op \ :: int \ int \ - bool"} and the locale instance is interpreted in the current - theory. *} - - interpretation %visible int: partial_order "op \ :: int \ int \ bool" -txt {* \normalsize - The argument of the command is a simple \emph{locale expression} - consisting of the name of the interpreted locale, which is - preceded by the qualifier @{text "int:"} and succeeded by a - white-space-separated list of terms, which provide a full - instantiation of the locale parameters. The parameters are referred - to by order of declaration, which is also the order in which - \isakeyword{print\_locale} outputs them. The locale has only a - single parameter, hence the list of instantiation terms is a - singleton. - - The command creates the goal - @{subgoals [display]} which can be shown easily: - *} - by unfold_locales auto - -text {* The effect of the command is that instances of all - conclusions of the locale are available in the theory, where names - are prefixed by the qualifier. For example, transitivity for @{typ - int} is named @{thm [source] int.trans} and is the following - theorem: - @{thm [display, indent=2] int.trans} - It is not possible to reference this theorem simply as @{text - trans}. This prevents unwanted hiding of existing theorems of the - theory by an interpretation. *} - - -subsection {* Second Version: Replacement of Definitions *} - -text {* Not only does the above interpretation qualify theorem names. - The prefix @{text int} is applied to all names introduced in locale - conclusions including names introduced in definitions. The - qualified name @{text int.less} is short for - the interpretation of the definition, which is @{term int.less}. - Qualified name and expanded form may be used almost - interchangeably.% -\footnote{Since @{term "op \"} is polymorphic, for @{term int.less} a - more general type will be inferred than for @{text int.less} which - is over type @{typ int}.} - The latter is preferred on output, as for example in the theorem - @{thm [source] int.less_le_trans}: @{thm [display, indent=2] - int.less_le_trans} - Both notations for the strict order are not satisfactory. The - constant @{term "op <"} is the strict order for @{typ int}. - In order to allow for the desired replacement, interpretation - accepts \emph{equations} in addition to the parameter instantiation. - These follow the locale expression and are indicated with the - keyword \isakeyword{where}. This is the revised interpretation: - *} -end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/Examples2.thy --- a/doc-src/Locales/Locales/Examples2.thy Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ -theory Examples2 -imports Examples -begin -text {* \vspace{-5ex} *} - interpretation %visible int: partial_order "op \ :: [int, int] \ bool" - where "int.less x y = (x < y)" - proof - - txt {* \normalsize The goals are now: - @{subgoals [display]} - The proof that~@{text \} is a partial order is as above. *} - show "partial_order (op \ :: int \ int \ bool)" - by unfold_locales auto - txt {* \normalsize The second goal is shown by unfolding the - definition of @{term "partial_order.less"}. *} - show "partial_order.less op \ x y = (x < y)" - unfolding partial_order.less_def [OF `partial_order op \`] - by auto - qed - -text {* Note that the above proof is not in the context of the - interpreted locale. Hence, the premise of @{text - "partial_order.less_def"} is discharged manually with @{text OF}. - *} -end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/Examples3.thy --- a/doc-src/Locales/Locales/Examples3.thy Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,659 +0,0 @@ -theory Examples3 -imports Examples -begin -text {* \vspace{-5ex} *} -subsection {* Third Version: Local Interpretation - \label{sec:local-interpretation} *} - -text {* In the above example, the fact that @{term "op \"} is a partial - order for the integers was used in the second goal to - discharge the premise in the definition of @{text "op \"}. In - general, proofs of the equations not only may involve definitions - from the interpreted locale but arbitrarily complex arguments in the - context of the locale. Therefore it would be convenient to have the - interpreted locale conclusions temporarily available in the proof. - This can be achieved by a locale interpretation in the proof body. - The command for local interpretations is \isakeyword{interpret}. We - repeat the example from the previous section to illustrate this. *} - - interpretation %visible int: partial_order "op \ :: int \ int \ bool" - where "int.less x y = (x < y)" - proof - - show "partial_order (op \ :: int \ int \ bool)" - by unfold_locales auto - then interpret int: partial_order "op \ :: [int, int] \ bool" . - show "int.less x y = (x < y)" - unfolding int.less_def by auto - qed - -text {* The inner interpretation is immediate from the preceding fact - and proved by assumption (Isar short hand ``.''). It enriches the - local proof context by the theorems - also obtained in the interpretation from Section~\ref{sec:po-first}, - and @{text int.less_def} may directly be used to unfold the - definition. Theorems from the local interpretation disappear after - leaving the proof context --- that is, after the succeeding - \isakeyword{next} or \isakeyword{qed} statement. *} - - -subsection {* Further Interpretations *} - -text {* Further interpretations are necessary for - the other locales. In @{text lattice} the operations~@{text \} - and~@{text \} are substituted by @{term "min :: int \ int \ int"} - and @{term "max :: int \ int \ int"}. The entire proof for the - interpretation is reproduced to give an example of a more - elaborate interpretation proof. Note that the equations are named - so they can be used in a later example. *} - - interpretation %visible int: lattice "op \ :: int \ int \ bool" - where int_min_eq: "int.meet x y = min x y" - and int_max_eq: "int.join x y = max x y" - proof - - show "lattice (op \ :: int \ int \ bool)" - txt {* \normalsize We have already shown that this is a partial - order, *} - apply unfold_locales - txt {* \normalsize hence only the lattice axioms remain to be - shown. - @{subgoals [display]} - By @{text is_inf} and @{text is_sup}, *} - apply (unfold int.is_inf_def int.is_sup_def) - txt {* \normalsize the goals are transformed to these - statements: - @{subgoals [display]} - This is Presburger arithmetic, which can be solved by the - method @{text arith}. *} - by arith+ - txt {* \normalsize In order to show the equations, we put ourselves - in a situation where the lattice theorems can be used in a - convenient way. *} - then interpret int: lattice "op \ :: int \ int \ bool" . - show "int.meet x y = min x y" - by (bestsimp simp: int.meet_def int.is_inf_def) - show "int.join x y = max x y" - by (bestsimp simp: int.join_def int.is_sup_def) - qed - -text {* Next follows that @{text "op \"} is a total order, again for - the integers. *} - - interpretation %visible int: total_order "op \ :: int \ int \ bool" - by unfold_locales arith - -text {* Theorems that are available in the theory at this point are shown in - Table~\ref{tab:int-lattice}. Two points are worth noting: - -\begin{table} -\hrule -\vspace{2ex} -\begin{center} -\begin{tabular}{l} - @{thm [source] int.less_def} from locale @{text partial_order}: \\ - \quad @{thm int.less_def} \\ - @{thm [source] int.meet_left} from locale @{text lattice}: \\ - \quad @{thm int.meet_left} \\ - @{thm [source] int.join_distr} from locale @{text distrib_lattice}: \\ - \quad @{thm int.join_distr} \\ - @{thm [source] int.less_total} from locale @{text total_order}: \\ - \quad @{thm int.less_total} -\end{tabular} -\end{center} -\hrule -\caption{Interpreted theorems for~@{text \} on the integers.} -\label{tab:int-lattice} -\end{table} - -\begin{itemize} -\item - Locale @{text distrib_lattice} was also interpreted. Since the - locale hierarchy reflects that total orders are distributive - lattices, the interpretation of the latter was inserted - automatically with the interpretation of the former. In general, - interpretation traverses the locale hierarchy upwards and interprets - all encountered locales, regardless whether imported or proved via - the \isakeyword{sublocale} command. Existing interpretations are - skipped avoiding duplicate work. -\item - The predicate @{term "op <"} appears in theorem @{thm [source] - int.less_total} - although an equation for the replacement of @{text "op \"} was only - given in the interpretation of @{text partial_order}. The - interpretation equations are pushed downwards the hierarchy for - related interpretations --- that is, for interpretations that share - the instances of parameters they have in common. -\end{itemize} - *} - -text {* The interpretations for a locale $n$ within the current - theory may be inspected with \isakeyword{print\_interps}~$n$. This - prints the list of instances of $n$, for which interpretations exist. - For example, \isakeyword{print\_interps} @{term partial_order} - outputs the following: -\begin{small} -\begin{alltt} - int! : partial_order "op \(\le\)" -\end{alltt} -\end{small} - Of course, there is only one interpretation. - The interpretation qualifier on the left is decorated with an - exclamation point. This means that it is mandatory. Qualifiers - can either be \emph{mandatory} or \emph{optional}, designated by - ``!'' or ``?'' respectively. Mandatory qualifiers must occur in a - name reference while optional ones need not. Mandatory qualifiers - prevent accidental hiding of names, while optional qualifiers can be - more convenient to use. For \isakeyword{interpretation}, the - default is ``!''. -*} - - -section {* Locale Expressions \label{sec:expressions} *} - -text {* - A map~@{term \} between partial orders~@{text \} and~@{text \} - is called order preserving if @{text "x \ y"} implies @{text "\ x \ - \ y"}. This situation is more complex than those encountered so - far: it involves two partial orders, and it is desirable to use the - existing locale for both. - - A locale for order preserving maps requires three parameters: @{text - le}~(\isakeyword{infixl}~@{text \}) and @{text - le'}~(\isakeyword{infixl}~@{text \}) for the orders and~@{text \} - for the map. - - In order to reuse the existing locale for partial orders, which has - the single parameter~@{text le}, it must be imported twice, once - mapping its parameter to~@{text le} from the new locale and once - to~@{text le'}. This can be achieved with a compound locale - expression. - - In general, a locale expression is a sequence of \emph{locale instances} - separated by~``$\textbf{+}$'' and followed by a \isakeyword{for} - clause. - An instance has the following format: -\begin{quote} - \textit{qualifier} \textbf{:} \textit{locale-name} - \textit{parameter-instantiation} -\end{quote} - We have already seen locale instances as arguments to - \isakeyword{interpretation} in Section~\ref{sec:interpretation}. - As before, the qualifier serves to disambiguate names from - different instances of the same locale. While in - \isakeyword{interpretation} qualifiers default to mandatory, in - import and in the \isakeyword{sublocale} command, they default to - optional. - - Since the parameters~@{text le} and~@{text le'} are to be partial - orders, our locale for order preserving maps will import the these - instances: -\begin{small} -\begin{alltt} - le: partial_order le - le': partial_order le' -\end{alltt} -\end{small} - For matter of convenience we choose to name parameter names and - qualifiers alike. This is an arbitrary decision. Technically, qualifiers - and parameters are unrelated. - - Having determined the instances, let us turn to the \isakeyword{for} - clause. It serves to declare locale parameters in the same way as - the context element \isakeyword{fixes} does. Context elements can - only occur after the import section, and therefore the parameters - referred to in the instances must be declared in the \isakeyword{for} - clause. The \isakeyword{for} clause is also where the syntax of these - parameters is declared. - - Two context elements for the map parameter~@{text \} and the - assumptions that it is order preserving complete the locale - declaration. *} - - locale order_preserving = - le: partial_order le + le': partial_order le' - for le (infixl "\" 50) and le' (infixl "\" 50) + - fixes \ - assumes hom_le: "x \ y \ \ x \ \ y" - -text (in order_preserving) {* Here are examples of theorems that are - available in the locale: - - \hspace*{1em}@{thm [source] hom_le}: @{thm hom_le} - - \hspace*{1em}@{thm [source] le.less_le_trans}: @{thm le.less_le_trans} - - \hspace*{1em}@{thm [source] le'.less_le_trans}: - @{thm [display, indent=4] le'.less_le_trans} - While there is infix syntax for the strict operation associated to - @{term "op \"}, there is none for the strict version of @{term - "op \"}. The abbreviation @{text less} with its infix syntax is only - available for the original instance it was declared for. We may - introduce the abbreviation @{text less'} with infix syntax~@{text \} - with the following declaration: *} - - abbreviation (in order_preserving) - less' (infixl "\" 50) where "less' \ partial_order.less le'" - -text (in order_preserving) {* Now the theorem is displayed nicely as - @{thm [source] le'.less_le_trans}: - @{thm [display, indent=2] le'.less_le_trans} *} - -text {* There are short notations for locale expressions. These are - discussed in the following. *} - - -subsection {* Default Instantiations *} - -text {* - It is possible to omit parameter instantiations. The - instantiation then defaults to the name of - the parameter itself. For example, the locale expression @{text - partial_order} is short for @{text "partial_order le"}, since the - locale's single parameter is~@{text le}. We took advantage of this - in the \isakeyword{sublocale} declarations of - Section~\ref{sec:changing-the-hierarchy}. *} - - -subsection {* Implicit Parameters \label{sec:implicit-parameters} *} - -text {* In a locale expression that occurs within a locale - declaration, omitted parameters additionally extend the (possibly - empty) \isakeyword{for} clause. - - The \isakeyword{for} clause is a general construct of Isabelle/Isar - to mark names occurring in the preceding declaration as ``arbitrary - but fixed''. This is necessary for example, if the name is already - bound in a surrounding context. In a locale expression, names - occurring in parameter instantiations should be bound by a - \isakeyword{for} clause whenever these names are not introduced - elsewhere in the context --- for example, on the left hand side of a - \isakeyword{sublocale} declaration. - - There is an exception to this rule in locale declarations, where the - \isakeyword{for} clause serves to declare locale parameters. Here, - locale parameters for which no parameter instantiation is given are - implicitly added, with their mixfix syntax, at the beginning of the - \isakeyword{for} clause. For example, in a locale declaration, the - expression @{text partial_order} is short for -\begin{small} -\begin{alltt} - partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.} -\end{alltt} -\end{small} - This short hand was used in the locale declarations throughout - Section~\ref{sec:import}. - *} - -text{* - The following locale declarations provide more examples. A - map~@{text \} is a lattice homomorphism if it preserves meet and - join. *} - - locale lattice_hom = - le: lattice + le': lattice le' for le' (infixl "\" 50) + - fixes \ - assumes hom_meet: "\ (x \ y) = le'.meet (\ x) (\ y)" - and hom_join: "\ (x \ y) = le'.join (\ x) (\ y)" - -text {* The parameter instantiation in the first instance of @{term - lattice} is omitted. This causes the parameter~@{text le} to be - added to the \isakeyword{for} clause, and the locale has - parameters~@{text le},~@{text le'} and, of course,~@{text \}. - - Before turning to the second example, we complete the locale by - providing infix syntax for the meet and join operations of the - second lattice. -*} - - context lattice_hom - begin - abbreviation meet' (infixl "\''" 50) where "meet' \ le'.meet" - abbreviation join' (infixl "\''" 50) where "join' \ le'.join" - end - -text {* The next example makes radical use of the short hand - facilities. A homomorphism is an endomorphism if both orders - coincide. *} - - locale lattice_end = lattice_hom _ le - -text {* The notation~@{text _} enables to omit a parameter in a - positional instantiation. The omitted parameter,~@{text le} becomes - the parameter of the declared locale and is, in the following - position, used to instantiate the second parameter of @{text - lattice_hom}. The effect is that of identifying the first in second - parameter of the homomorphism locale. *} - -text {* The inheritance diagram of the situation we have now is shown - in Figure~\ref{fig:hom}, where the dashed line depicts an - interpretation which is introduced below. Parameter instantiations - are indicated by $\sqsubseteq \mapsto \preceq$ etc. By looking at - the inheritance diagram it would seem - that two identical copies of each of the locales @{text - partial_order} and @{text lattice} are imported by @{text - lattice_end}. This is not the case! Inheritance paths with - identical morphisms are automatically detected and - the conclusions of the respective locales appear only once. - -\begin{figure} -\hrule \vspace{2ex} -\begin{center} -\begin{tikzpicture} - \node (o) at (0,0) {@{text partial_order}}; - \node (oh) at (1.5,-2) {@{text order_preserving}}; - \node (oh1) at (1.5,-0.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; - \node (oh2) at (0,-1.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; - \node (l) at (-1.5,-2) {@{text lattice}}; - \node (lh) at (0,-4) {@{text lattice_hom}}; - \node (lh1) at (0,-2.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; - \node (lh2) at (-1.5,-3.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; - \node (le) at (0,-6) {@{text lattice_end}}; - \node (le1) at (0,-4.8) - [anchor=west]{$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; - \node (le2) at (0,-5.2) - [anchor=west]{$\scriptscriptstyle \preceq \mapsto \sqsubseteq$}; - \draw (o) -- (l); - \draw[dashed] (oh) -- (lh); - \draw (lh) -- (le); - \draw (o) .. controls (oh1.south west) .. (oh); - \draw (o) .. controls (oh2.north east) .. (oh); - \draw (l) .. controls (lh1.south west) .. (lh); - \draw (l) .. controls (lh2.north east) .. (lh); -\end{tikzpicture} -\end{center} -\hrule -\caption{Hierarchy of Homomorphism Locales.} -\label{fig:hom} -\end{figure} - *} - -text {* It can be shown easily that a lattice homomorphism is order - preserving. As the final example of this section, a locale - interpretation is used to assert this: *} - - sublocale lattice_hom \ order_preserving - proof unfold_locales - fix x y - assume "x \ y" - then have "y = (x \ y)" by (simp add: le.join_connection) - then have "\ y = (\ x \' \ y)" by (simp add: hom_join [symmetric]) - then show "\ x \ \ y" by (simp add: le'.join_connection) - qed - -text (in lattice_hom) {* - Theorems and other declarations --- syntax, in particular --- from - the locale @{text order_preserving} are now active in @{text - lattice_hom}, for example - @{thm [source] hom_le}: - @{thm [display, indent=2] hom_le} - This theorem will be useful in the following section. - *} - - -section {* Conditional Interpretation *} - -text {* There are situations where an interpretation is not possible - in the general case since the desired property is only valid if - certain conditions are fulfilled. Take, for example, the function - @{text "\i. n * i"} that scales its argument by a constant factor. - This function is order preserving (and even a lattice endomorphism) - with respect to @{term "op \"} provided @{text "n \ 0"}. - - It is not possible to express this using a global interpretation, - because it is in general unspecified whether~@{term n} is - non-negative, but one may make an interpretation in an inner context - of a proof where full information is available. - This is not fully satisfactory either, since potentially - interpretations may be required to make interpretations in many - contexts. What is - required is an interpretation that depends on the condition --- and - this can be done with the \isakeyword{sublocale} command. For this - purpose, we introduce a locale for the condition. *} - - locale non_negative = - fixes n :: int - assumes non_neg: "0 \ n" - -text {* It is again convenient to make the interpretation in an - incremental fashion, first for order preserving maps, the for - lattice endomorphisms. *} - - sublocale non_negative \ - order_preserving "op \" "op \" "\i. n * i" - using non_neg by unfold_locales (rule mult_left_mono) - -text {* While the proof of the previous interpretation - is straightforward from monotonicity lemmas for~@{term "op *"}, the - second proof follows a useful pattern. *} - - sublocale %visible non_negative \ lattice_end "op \" "\i. n * i" - proof (unfold_locales, unfold int_min_eq int_max_eq) - txt {* \normalsize Unfolding the locale predicates \emph{and} the - interpretation equations immediately yields two subgoals that - reflect the core conjecture. - @{subgoals [display]} - It is now necessary to show, in the context of @{term - non_negative}, that multiplication by~@{term n} commutes with - @{term min} and @{term max}. *} - qed (auto simp: hom_le) - -text (in order_preserving) {* The lemma @{thm [source] hom_le} - simplifies a proof that would have otherwise been lengthy and we may - consider making it a default rule for the simplifier: *} - - lemmas (in order_preserving) hom_le [simp] - - -subsection {* Avoiding Infinite Chains of Interpretations - \label{sec:infinite-chains} *} - -text {* Similar situations arise frequently in formalisations of - abstract algebra where it is desirable to express that certain - constructions preserve certain properties. For example, polynomials - over rings are rings, or --- an example from the domain where the - illustrations of this tutorial are taken from --- a partial order - may be obtained for a function space by point-wise lifting of the - partial order of the co-domain. This corresponds to the following - interpretation: *} - - sublocale %visible partial_order \ f: partial_order "\f g. \x. f x \ g x" - oops - -text {* Unfortunately this is a cyclic interpretation that leads to an - infinite chain, namely - @{text [display, indent=2] "partial_order \ partial_order (\f g. \x. f x \ g x) \ - partial_order (\f g. \x y. f x y \ g x y) \ \"} - and the interpretation is rejected. - - Instead it is necessary to declare a locale that is logically - equivalent to @{term partial_order} but serves to collect facts - about functions spaces where the co-domain is a partial order, and - to make the interpretation in its context: *} - - locale fun_partial_order = partial_order - - sublocale fun_partial_order \ - f: partial_order "\f g. \x. f x \ g x" - by unfold_locales (fast,rule,fast,blast intro: trans) - -text {* It is quite common in abstract algebra that such a construction - maps a hierarchy of algebraic structures (or specifications) to a - related hierarchy. By means of the same lifting, a function space - is a lattice if its co-domain is a lattice: *} - - locale fun_lattice = fun_partial_order + lattice - - sublocale fun_lattice \ f: lattice "\f g. \x. f x \ g x" - proof unfold_locales - fix f g - have "partial_order.is_inf (\f g. \x. f x \ g x) f g (\x. f x \ g x)" - apply (rule is_infI) apply rule+ apply (drule spec, assumption)+ done - then show "\inf. partial_order.is_inf (\f g. \x. f x \ g x) f g inf" - by fast - next - fix f g - have "partial_order.is_sup (\f g. \x. f x \ g x) f g (\x. f x \ g x)" - apply (rule is_supI) apply rule+ apply (drule spec, assumption)+ done - then show "\sup. partial_order.is_sup (\f g. \x. f x \ g x) f g sup" - by fast - qed - - -section {* Further Reading *} - -text {* More information on locales and their interpretation is - available. For the locale hierarchy of import and interpretation - dependencies see~\cite{Ballarin2006a}; interpretations in theories - and proofs are covered in~\cite{Ballarin2006b}. In the latter, I - show how interpretation in proofs enables to reason about families - of algebraic structures, which cannot be expressed with locales - directly. - - Haftmann and Wenzel~\cite{HaftmannWenzel2007} overcome a restriction - of axiomatic type classes through a combination with locale - interpretation. The result is a Haskell-style class system with a - facility to generate ML and Haskell code. Classes are sufficient for - simple specifications with a single type parameter. The locales for - orders and lattices presented in this tutorial fall into this - category. Order preserving maps, homomorphisms and vector spaces, - on the other hand, do not. - - The locales reimplementation for Isabelle 2009 provides, among other - improvements, a clean integration with Isabelle/Isar's local theory - mechanisms, which are described in another paper by Haftmann and - Wenzel~\cite{HaftmannWenzel2009}. - - The original work of Kamm\"uller on locales~\cite{KammullerEtAl1999} - may be of interest from a historical perspective. My previous - report on locales and locale expressions~\cite{Ballarin2004a} - describes a simpler form of expressions than available now and is - outdated. The mathematical background on orders and lattices is - taken from Jacobson's textbook on algebra~\cite[Chapter~8]{Jacobson1985}. - - The sources of this tutorial, which include all proofs, are - available with the Isabelle distribution at - \url{http://isabelle.in.tum.de}. - *} - -text {* -\begin{table} -\hrule -\vspace{2ex} -\begin{center} -\begin{tabular}{l>$c<$l} - \multicolumn{3}{l}{Miscellaneous} \\ - - \textit{attr-name} & ::= - & \textit{name} $|$ \textit{attribute} $|$ - \textit{name} \textit{attribute} \\ - \textit{qualifier} & ::= - & \textit{name} [``\textbf{?}'' $|$ ``\textbf{!}''] \\[2ex] - - \multicolumn{3}{l}{Context Elements} \\ - - \textit{fixes} & ::= - & \textit{name} [ ``\textbf{::}'' \textit{type} ] - [ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ - \textit{mixfix} ] \\ -\begin{comment} - \textit{constrains} & ::= - & \textit{name} ``\textbf{::}'' \textit{type} \\ -\end{comment} - \textit{assumes} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ -\begin{comment} - \textit{defines} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ - \textit{notes} & ::= - & [ \textit{attr-name} ``\textbf{=}'' ] - ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ -\end{comment} - - \textit{element} & ::= - & \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ -\begin{comment} - & | - & \textbf{constrains} \textit{constrains} - ( \textbf{and} \textit{constrains} )$^*$ \\ -\end{comment} - & | - & \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\[2ex] -%\begin{comment} -% & | -% & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ -% & | -% & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ -%\end{comment} - - \multicolumn{3}{l}{Locale Expressions} \\ - - \textit{pos-insts} & ::= - & ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\ - \textit{named-insts} & ::= - & \textbf{where} \textit{name} ``\textbf{=}'' \textit{term} - ( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\ - \textit{instance} & ::= - & [ \textit{qualifier} ``\textbf{:}'' ] - \textit{name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\ - \textit{expression} & ::= - & \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$ - [ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex] - - \multicolumn{3}{l}{Declaration of Locales} \\ - - \textit{locale} & ::= - & \textit{element}$^+$ \\ - & | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ - \textit{toplevel} & ::= - & \textbf{locale} \textit{name} [ ``\textbf{=}'' - \textit{locale} ] \\[2ex] - - \multicolumn{3}{l}{Interpretation} \\ - - \textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] - \textit{prop} \\ - \textit{equations} & ::= & \textbf{where} \textit{equation} ( \textbf{and} - \textit{equation} )$^*$ \\ - \textit{toplevel} & ::= - & \textbf{sublocale} \textit{name} ( ``$<$'' $|$ - ``$\subseteq$'' ) \textit{expression} \textit{proof} \\ - & | - & \textbf{interpretation} - \textit{expression} [ \textit{equations} ] \textit{proof} \\ - & | - & \textbf{interpret} - \textit{expression} \textit{proof} \\[2ex] - - \multicolumn{3}{l}{Diagnostics} \\ - - \textit{toplevel} & ::= - & \textbf{print\_locales} \\ - & | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{name} \\ - & | & \textbf{print\_interps} \textit{name} -\end{tabular} -\end{center} -\hrule -\caption{Syntax of Locale Commands.} -\label{tab:commands} -\end{table} - *} - -text {* \textbf{Revision History.} For the present third revision of - the tutorial, much of the explanatory text - was rewritten. Inheritance of interpretation equations is - available with the forthcoming release of Isabelle, which at the - time of editing these notes is expected for the end of 2009. - The second revision accommodates changes introduced by the locales - reimplementation for Isabelle 2009. Most notably locale expressions - have been generalised from renaming to instantiation. *} - -text {* \textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, - Randy Pollack, Andreas Schropp, Christian Sternagel and Makarius Wenzel - have made - useful comments on earlier versions of this document. The section - on conditional interpretation was inspired by a number of e-mail - enquiries the author received from locale users, and which suggested - that this use case is important enough to deserve explicit - explanation. The term \emph{conditional interpretation} is due to - Larry Paulson. *} - -end diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/Examples.tex --- a/doc-src/Locales/Locales/document/Examples.tex Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1520 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Examples}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Examples\isanewline -\isakeyword{imports}\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -% -\isataginvisible -\isacommand{pretty{\isaliteral{5F}{\isacharunderscore}}setmargin}\isamarkupfalse% -\ {\isadigit{6}}{\isadigit{5}}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsection{Introduction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locales are based on contexts. A \emph{context} can be seen as a - formula schema -\[ - \isa{{\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5C3C646F74733E}{\isasymdots}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ A\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3B}{\isacharsemicolon}}A\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}} -\] - where the variables~\isa{x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}, \ldots,~\isa{x\isaliteral{5C3C5E7375623E}{}\isactrlsub n} are called - \emph{parameters} and the premises $\isa{A\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}, \ldots,~\isa{A\isaliteral{5C3C5E7375623E}{}\isactrlsub m}$ \emph{assumptions}. A formula~\isa{C} - is a \emph{theorem} in the context if it is a conclusion -\[ - \isa{{\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5C3C646F74733E}{\isasymdots}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ A\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3B}{\isacharsemicolon}}A\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C}. -\] - Isabelle/Isar's notion of context goes beyond this logical view. - Its contexts record, in a consecutive order, proved - conclusions along with \emph{attributes}, which can provide context - specific configuration information for proof procedures and concrete - syntax. From a logical perspective, locales are just contexts that - have been made persistent. To the user, though, they provide - powerful means for declaring and combining contexts, and for the - reuse of theorems proved in these contexts.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Simple Locales% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In its simplest form, a - \emph{locale declaration} consists of a sequence of context elements - declaring parameters (keyword \isakeyword{fixes}) and assumptions - (keyword \isakeyword{assumes}). The following is the specification of - partial orders, as locale \isa{partial{\isaliteral{5F}{\isacharunderscore}}order}.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ \isakeyword{fixes}\ le\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isakeyword{assumes}\ refl\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ anti{\isaliteral{5F}{\isacharunderscore}}sym\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ trans\ {\isaliteral{5B}{\isacharbrackleft}}trans{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -The parameter of this locale is~\isa{le}, - which is a binary predicate with infix syntax~\isa{{\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}}. The - parameter syntax is available in the subsequent - assumptions, which are the familiar partial order axioms. - - Isabelle recognises unbound names as free variables. In locale - assumptions, these are implicitly universally quantified. That is, - \isa{{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z} in fact means - \isa{{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y\ z{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z}. - - Two commands are provided to inspect locales: - \isakeyword{print\_locales} lists the names of all locales of the - current theory; \isakeyword{print\_locale}~$n$ prints the parameters - and assumptions of locale $n$; the variation \isakeyword{print\_locale!}~$n$ - additionally outputs the conclusions that are stored in the locale. - We may inspect the new locale - by issuing \isakeyword{print\_locale!} \isa{partial{\isaliteral{5F}{\isacharunderscore}}order}. The output - is the following list of context elements. -\begin{small} -\begin{alltt} - \isakeyword{fixes} le :: "'a \(\Rightarrow\) 'a \(\Rightarrow\) bool" (\isakeyword{infixl} "\(\sqsubseteq\)" 50) - \isakeyword{assumes} "partial_order op \(\sqsubseteq\)" - \isakeyword{notes} assumption - refl [intro, simp] = `?x \(\sqsubseteq\) ?x` - \isakeyword{and} - anti_sym [intro] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?x\(\isasymrbrakk\) \(\Longrightarrow\) ?x = ?y` - \isakeyword{and} - trans [trans] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?z\(\isasymrbrakk\) \(\Longrightarrow\) ?x \(\sqsubseteq\) ?z` -\end{alltt} -\end{small} - The keyword \isakeyword{notes} denotes a conclusion element. There - is one conclusion, which was added automatically. Instead, there is - only one assumption, namely \isa{partial{\isaliteral{5F}{\isacharunderscore}}order\ op\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}}. The locale - declaration has introduced the predicate \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} to the theory. This predicate is the - \emph{locale predicate}. Its definition may be inspected by - issuing \isakeyword{thm} \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{5F}{\isacharunderscore}}def}. - \begin{isabelle}% -\ \ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\isanewline -\isaindent{\ \ }{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3F}{\isacharquery}}le\ x\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline -\isaindent{\ \ }{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3F}{\isacharquery}}le\ x\ y\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}le\ y\ x\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline -\isaindent{\ \ }{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x\ y\ z{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3F}{\isacharquery}}le\ x\ y\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}le\ y\ z\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}le\ x\ z{\isaliteral{29}{\isacharparenright}}% -\end{isabelle} - In our example, this is a unary predicate over the parameter of the - locale. It is equivalent to the original assumptions, which have - been turned into conclusions and are - available as theorems in the context of the locale. The names and - attributes from the locale declaration are associated to these - theorems and are effective in the context of the locale. - - Each conclusion has a \emph{foundational theorem} as counterpart - in the theory. Technically, this is simply the theorem composed - of context and conclusion. For the transitivity theorem, this is - \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}trans}: - \begin{isabelle}% -\ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{3F}{\isacharquery}}le{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}z% -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Targets: Extending Locales% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The specification of a locale is fixed, but its list of conclusions - may be extended through Isar commands that take a \emph{target} argument. - In the following, \isakeyword{definition} and - \isakeyword{theorem} are illustrated. - Table~\ref{tab:commands-with-target} lists Isar commands that accept - a target. Isar provides various ways of specifying the target. A - target for a single command may be indicated with keyword - \isakeyword{in} in the following way: - -\begin{table} -\hrule -\vspace{2ex} -\begin{center} -\begin{tabular}{ll} - \isakeyword{definition} & definition through an equation \\ - \isakeyword{inductive} & inductive definition \\ - \isakeyword{primrec} & primitive recursion \\ - \isakeyword{fun}, \isakeyword{function} & general recursion \\ - \isakeyword{abbreviation} & syntactic abbreviation \\ - \isakeyword{theorem}, etc.\ & theorem statement with proof \\ - \isakeyword{theorems}, etc.\ & redeclaration of theorems \\ - \isakeyword{text}, etc.\ & document markup -\end{tabular} -\end{center} -\hrule -\caption{Isar commands that accept a target.} -\label{tab:commands-with-target} -\end{table}% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{definition}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ less\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C616E643E}{\isasymand}}\ x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -The strict order \isa{less} with infix - syntax~\isa{{\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}} is - defined in terms of the locale parameter~\isa{le} and the general - equality of the object logic we work in. The definition generates a - \emph{foundational constant} - \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less} with definition \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}def}: - \begin{isabelle}% -\ \ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\isaindent{\ \ }partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ {\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}le\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{29}{\isacharparenright}}% -\end{isabelle} - At the same time, the locale is extended by syntax transformations - hiding this construction in the context of the locale. Here, the - abbreviation \isa{less} is available for - \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ le}, and it is printed - and parsed as infix~\isa{{\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}}. Finally, the conclusion \isa{less{\isaliteral{5F}{\isacharunderscore}}def} is added to the locale: - \begin{isabelle}% -\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{29}{\isacharparenright}}% -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The treatment of theorem statements is more straightforward. - As an example, here is the derivation of a transitivity law for the - strict order relation.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{29}{\isacharparenright}}\ less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans\ {\isaliteral{5B}{\isacharbrackleft}}trans{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ y{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimvisible -\ \ \ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{unfolding}\isamarkupfalse% -\ less{\isaliteral{5F}{\isacharunderscore}}def\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}blast\ intro{\isaliteral{3A}{\isacharcolon}}\ trans{\isaliteral{29}{\isacharparenright}}% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -In the context of the proof, conclusions of the - locale may be used like theorems. Attributes are effective: \isa{anti{\isaliteral{5F}{\isacharunderscore}}sym} was - declared as introduction rule, hence it is in the context's set of - rules used by the classical reasoner by default.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Context Blocks% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When working with locales, sequences of commands with the same - target are frequent. A block of commands, delimited by - \isakeyword{begin} and \isakeyword{end}, makes a theory-like style - of working possible. All commands inside the block refer to the - same target. A block may immediately follow a locale - declaration, which makes that locale the target. Alternatively the - target for a block may be given with the \isakeyword{context} - command. - - This style of working is illustrated in the block below, where - notions of infimum and supremum for partial orders are introduced, - together with theorems about their uniqueness.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{context}\isamarkupfalse% -\ partial{\isaliteral{5F}{\isacharunderscore}}order\isanewline -\ \ \isakeyword{begin}\isanewline -\isanewline -\ \ \isacommand{definition}\isamarkupfalse% -\isanewline -\ \ \ \ is{\isaliteral{5F}{\isacharunderscore}}inf\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C616E643E}{\isasymand}}\ i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C616E643E}{\isasymand}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ i{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{definition}\isamarkupfalse% -\isanewline -\ \ \ \ is{\isaliteral{5F}{\isacharunderscore}}sup\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C616E643E}{\isasymand}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C616E643E}{\isasymand}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ s\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadeliminvisible -\isanewline -\ \ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}infI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}z{\isaliteral{2E}{\isachardot}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}lower\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}greatest\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}uniq{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{3B}{\isacharsemicolon}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{27}{\isacharprime}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ i\ {\isaliteral{3D}{\isacharequal}}\ i{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ inf{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ inf{\isaliteral{27}{\isacharprime}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ anti{\isaliteral{5F}{\isacharunderscore}}sym{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ inf{\isaliteral{27}{\isacharprime}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ i{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}greatest{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ inf\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ inf\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ inf\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}i{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}greatest{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ inf{\isaliteral{27}{\isacharprime}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}i{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ inf{\isaliteral{27}{\isacharprime}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}i{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -% -\isadeliminvisible -\isanewline -\ \ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{theorem}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}related\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}supI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}z{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ s\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}least\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ s\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}upper\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}uniq{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{3B}{\isacharsemicolon}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{27}{\isacharprime}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ s\ {\isaliteral{3D}{\isacharequal}}\ s{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ sup{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ sup{\isaliteral{27}{\isacharprime}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ anti{\isaliteral{5F}{\isacharunderscore}}sym{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ sup\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}s\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}least{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ sup{\isaliteral{27}{\isacharprime}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ sup{\isaliteral{27}{\isacharprime}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ sup{\isaliteral{27}{\isacharprime}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}s{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}least{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ sup\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ sup\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -% -\isadeliminvisible -\isanewline -\ \ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{theorem}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}related\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -\isanewline -\ \ \isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -The syntax of the locale commands discussed in this tutorial is - shown in Table~\ref{tab:commands}. The grammar is complete with the - exception of the context elements \isakeyword{constrains} and - \isakeyword{defines}, which are provided for backward - compatibility. See the Isabelle/Isar Reference - Manual~\cite{IsarRef} for full documentation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Import \label{sec:import}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Algebraic structures are commonly defined by adding operations and - properties to existing structures. For example, partial orders - are extended to lattices and total orders. Lattices are extended to - distributive lattices.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -With locales, this kind of inheritance is achieved through - \emph{import} of locales. The import part of a locale declaration, - if present, precedes the context elements. Here is an example, - where partial orders are extended to lattices.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ lattice\ {\isaliteral{3D}{\isacharequal}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{2B}{\isacharplus}}\isanewline -\ \ \ \ \isakeyword{assumes}\ ex{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}inf{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ inf{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ ex{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}sup{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ sup{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{begin}% -\begin{isamarkuptext}% -These assumptions refer to the predicates for infimum - and supremum defined for \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} in the previous - section. We now introduce the notions of meet and join.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{definition}\isamarkupfalse% -\isanewline -\ \ \ \ meet\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{7}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}THE\ inf{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ inf{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{definition}\isamarkupfalse% -\isanewline -\ \ \ \ join\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{6}}{\isadigit{5}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}THE\ sup{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ sup{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadeliminvisible -\isanewline -\ \ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{lemma}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}equality\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ meet{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}THE\ i{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ the{\isaliteral{5F}{\isacharunderscore}}equality{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}uniq\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{60}{\isacharbackquoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ meetI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ i\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}z{\isaliteral{2E}{\isachardot}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ meet{\isaliteral{5F}{\isacharunderscore}}equality{\isaliteral{2C}{\isacharcomma}}\ rule\ is{\isaliteral{5F}{\isacharunderscore}}infI{\isaliteral{29}{\isacharparenright}}\ blast{\isaliteral{2B}{\isacharplus}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}meet\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ meet{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ ex{\isaliteral{5F}{\isacharunderscore}}inf\ \isacommand{obtain}\isamarkupfalse% -\ i\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ {\isaliteral{28}{\isacharparenleft}}THE\ i{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ theI{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}uniq\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{60}{\isacharbackquoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ i{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}left\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}lower{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}right\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}lower{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}le\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{3B}{\isacharsemicolon}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}greatest{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}equality\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ join{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}THE\ s{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ the{\isaliteral{5F}{\isacharunderscore}}equality{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}uniq\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{60}{\isacharbackquoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ joinI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ s\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}z{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ s\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ join{\isaliteral{5F}{\isacharunderscore}}equality{\isaliteral{2C}{\isacharcomma}}\ rule\ is{\isaliteral{5F}{\isacharunderscore}}supI{\isaliteral{29}{\isacharparenright}}\ blast{\isaliteral{2B}{\isacharplus}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}join\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ join{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ ex{\isaliteral{5F}{\isacharunderscore}}sup\ \isacommand{obtain}\isamarkupfalse% -\ s\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ {\isaliteral{28}{\isacharparenleft}}THE\ s{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ theI{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}uniq\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{60}{\isacharbackquoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ s{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}left\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}upper{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}right\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}upper{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}le\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}least{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}assoc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ meetI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ w\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}commute{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ meetI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{5F}{\isacharunderscore}}absorb{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ meetI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}assoc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ joinI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ w\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ w{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}commute{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ joinI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{5F}{\isacharunderscore}}absorb{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ joinI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}idem{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ meet{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{5F}{\isacharunderscore}}absorb{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ join{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{5F}{\isacharunderscore}}absorb{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}related\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ meetI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}drule\ meet{\isaliteral{5F}{\isacharunderscore}}related{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ meet{\isaliteral{5F}{\isacharunderscore}}commute{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}related\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ joinI{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ z\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}drule\ join{\isaliteral{5F}{\isacharunderscore}}related{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ join{\isaliteral{5F}{\isacharunderscore}}commute{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}connection{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}connection{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}connection{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{using}\isamarkupfalse% -\ meet{\isaliteral{5F}{\isacharunderscore}}commute\ meet{\isaliteral{5F}{\isacharunderscore}}connection\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isanewline -\ \ \isacommand{theorem}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}connection{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{using}\isamarkupfalse% -\ join{\isaliteral{5F}{\isacharunderscore}}commute\ join{\isaliteral{5F}{\isacharunderscore}}connection\ \isacommand{by}\isamarkupfalse% -\ simp% -\begin{isamarkuptext}% -Naming according to Jacobson I, p.\ 459.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemmas}\isamarkupfalse% -\ L{\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ join{\isaliteral{5F}{\isacharunderscore}}commute\ meet{\isaliteral{5F}{\isacharunderscore}}commute\isanewline -\ \ \isacommand{lemmas}\isamarkupfalse% -\ L{\isadigit{2}}\ {\isaliteral{3D}{\isacharequal}}\ join{\isaliteral{5F}{\isacharunderscore}}assoc\ meet{\isaliteral{5F}{\isacharunderscore}}assoc\isanewline -\ \ \isanewline -\ \ \isacommand{lemmas}\isamarkupfalse% -\ L{\isadigit{4}}\ {\isaliteral{3D}{\isacharequal}}\ join{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{5F}{\isacharunderscore}}absorb\ meet{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{5F}{\isacharunderscore}}absorb% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -\isanewline -\ \ \isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Locales for total orders and distributive lattices follow to - establish a sufficiently rich landscape of locales for - further examples in this tutorial. Each comes with an example - theorem.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ total{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{3D}{\isacharequal}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{2B}{\isacharplus}}\isanewline -\ \ \ \ \isakeyword{assumes}\ total{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C6F723E}{\isasymor}}\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ total{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{29}{\isacharparenright}}\ less{\isaliteral{5F}{\isacharunderscore}}total{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ y\ {\isaliteral{5C3C6F723E}{\isasymor}}\ x\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C6F723E}{\isasymor}}\ y\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{using}\isamarkupfalse% -\ total\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ less{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\ \ \isacommand{locale}\isamarkupfalse% -\ distrib{\isaliteral{5F}{\isacharunderscore}}lattice\ {\isaliteral{3D}{\isacharequal}}\ lattice\ {\isaliteral{2B}{\isacharplus}}\isanewline -\ \ \ \ \isakeyword{assumes}\ meet{\isaliteral{5F}{\isacharunderscore}}distr{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ distrib{\isaliteral{5F}{\isacharunderscore}}lattice{\isaliteral{29}{\isacharparenright}}\ join{\isaliteral{5F}{\isacharunderscore}}distr{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \ \isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ L{\isadigit{4}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ L{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ L{\isadigit{1}}\ meet{\isaliteral{5F}{\isacharunderscore}}distr{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ L{\isadigit{1}}\ L{\isadigit{4}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ meet{\isaliteral{5F}{\isacharunderscore}}distr{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -The locale hierarchy obtained through these declarations is shown in - Figure~\ref{fig:lattices}(a). - -\begin{figure} -\hrule \vspace{2ex} -\begin{center} -\subfigure[Declared hierarchy]{ -\begin{tikzpicture} - \node (po) at (0,0) {\isa{partial{\isaliteral{5F}{\isacharunderscore}}order}}; - \node (lat) at (-1.5,-1) {\isa{lattice}}; - \node (dlat) at (-1.5,-2) {\isa{distrib{\isaliteral{5F}{\isacharunderscore}}lattice}}; - \node (to) at (1.5,-1) {\isa{total{\isaliteral{5F}{\isacharunderscore}}order}}; - \draw (po) -- (lat); - \draw (lat) -- (dlat); - \draw (po) -- (to); -% \draw[->, dashed] (lat) -- (to); -\end{tikzpicture} -} \\ -\subfigure[Total orders are lattices]{ -\begin{tikzpicture} - \node (po) at (0,0) {\isa{partial{\isaliteral{5F}{\isacharunderscore}}order}}; - \node (lat) at (0,-1) {\isa{lattice}}; - \node (dlat) at (-1.5,-2) {\isa{distrib{\isaliteral{5F}{\isacharunderscore}}lattice}}; - \node (to) at (1.5,-2) {\isa{total{\isaliteral{5F}{\isacharunderscore}}order}}; - \draw (po) -- (lat); - \draw (lat) -- (dlat); - \draw (lat) -- (to); -% \draw[->, dashed] (dlat) -- (to); -\end{tikzpicture} -} \quad -\subfigure[Total orders are distributive lattices]{ -\begin{tikzpicture} - \node (po) at (0,0) {\isa{partial{\isaliteral{5F}{\isacharunderscore}}order}}; - \node (lat) at (0,-1) {\isa{lattice}}; - \node (dlat) at (0,-2) {\isa{distrib{\isaliteral{5F}{\isacharunderscore}}lattice}}; - \node (to) at (0,-3) {\isa{total{\isaliteral{5F}{\isacharunderscore}}order}}; - \draw (po) -- (lat); - \draw (lat) -- (dlat); - \draw (dlat) -- (to); -\end{tikzpicture} -} -\end{center} -\hrule -\caption{Hierarchy of Lattice Locales.} -\label{fig:lattices} -\end{figure}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Changing the Locale Hierarchy - \label{sec:changing-the-hierarchy}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locales enable to prove theorems abstractly, relative to - sets of assumptions. These theorems can then be used in other - contexts where the assumptions themselves, or - instances of the assumptions, are theorems. This form of theorem - reuse is called \emph{interpretation}. Locales generalise - interpretation from theorems to conclusions, enabling the reuse of - definitions and other constructs that are not part of the - specifications of the locales. - - The first form of interpretation we will consider in this tutorial - is provided by the \isakeyword{sublocale} command. It enables to - modify the import hierarchy to reflect the \emph{logical} relation - between locales. - - Consider the locale hierarchy from Figure~\ref{fig:lattices}(a). - Total orders are lattices, although this is not reflected here, and - definitions, theorems and other conclusions - from \isa{lattice} are not available in \isa{total{\isaliteral{5F}{\isacharunderscore}}order}. To - obtain the situation in Figure~\ref{fig:lattices}(b), it is - sufficient to add the conclusions of the latter locale to the former. - The \isakeyword{sublocale} command does exactly this. - The declaration \isakeyword{sublocale} $l_1 - \subseteq l_2$ causes locale $l_2$ to be \emph{interpreted} in the - context of $l_1$. This means that all conclusions of $l_2$ are made - available in $l_1$. - - Of course, the change of hierarchy must be supported by a theorem - that reflects, in our example, that total orders are indeed - lattices. Therefore the \isakeyword{sublocale} command generates a - goal, which must be discharged by the user. This is illustrated in - the following paragraphs. First the sublocale relation is stated.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{sublocale}\isamarkupfalse% -\ total{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ lattice% -\begin{isamarkuptxt}% -\normalsize - This enters the context of locale \isa{total{\isaliteral{5F}{\isacharunderscore}}order}, in - which the goal \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ lattice\ op\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}% -\end{isabelle} must be shown. - Now the - locale predicate needs to be unfolded --- for example, using its - definition or by introduction rules - provided by the locale package. For automation, the locale package - provides the methods \isa{intro{\isaliteral{5F}{\isacharunderscore}}locales} and \isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}. They are aware of the - current context and dependencies between locales and automatically - discharge goals implied by these. While \isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales} - always unfolds locale predicates to assumptions, \isa{intro{\isaliteral{5F}{\isacharunderscore}}locales} only unfolds definitions along the locale - hierarchy, leaving a goal consisting of predicates defined by the - locale package. Occasionally the latter is of advantage since the goal - is smaller. - - For the current goal, we would like to get hold of - the assumptions of \isa{lattice}, which need to be shown, hence - \isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales} is appropriate.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{proof}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales% -\begin{isamarkuptxt}% -\normalsize - Since the fact that both lattices and total orders are partial - orders is already reflected in the locale hierarchy, the assumptions - of \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} are discharged automatically, and only the - assumptions introduced in \isa{lattice} remain as subgoals - \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}inf{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ inf\isanewline -\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}sup{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ sup% -\end{isabelle} - The proof for the first subgoal is obtained by constructing an - infimum, whose existence is implied by totality.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ total\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ {\isaliteral{28}{\isacharparenleft}}if\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ then\ x\ else\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}auto\ simp{\isaliteral{3A}{\isacharcolon}}\ is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}inf{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}inf\ x\ y\ inf{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\begin{isamarkuptxt}% -\normalsize - The proof for the second subgoal is analogous and not - reproduced here.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\isadeliminvisible -\ \ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ total\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ {\isaliteral{28}{\isacharparenleft}}if\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ then\ y\ else\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}auto\ simp{\isaliteral{3A}{\isacharcolon}}\ is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}sup{\isaliteral{2E}{\isachardot}}\ is{\isaliteral{5F}{\isacharunderscore}}sup\ x\ y\ sup{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isadelimvisible -\ % -\endisadelimvisible -% -\isatagvisible -\isacommand{qed}\isamarkupfalse% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -Similarly, we may establish that total orders are distributive - lattices with a second \isakeyword{sublocale} statement.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{sublocale}\isamarkupfalse% -\ total{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ distrib{\isaliteral{5F}{\isacharunderscore}}lattice\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\ z\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{is}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}l\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}r{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}% -\begin{isamarkuptxt}% -Jacobson I, p.\ 462% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ c{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}z\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ c\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}l\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}metis\ c\ join{\isaliteral{5F}{\isacharunderscore}}connection{\isadigit{2}}\ join{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ meet{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ total{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{from}\isamarkupfalse% -\ c\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}r{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}metis\ meet{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}l\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}r{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ c{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C6F723E}{\isasymor}}\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% -\ c\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}l\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}metis\ join{\isaliteral{5F}{\isacharunderscore}}connection{\isadigit{2}}\ join{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ meet{\isaliteral{5F}{\isacharunderscore}}connection\ total\ trans{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{from}\isamarkupfalse% -\ c\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}r{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}metis\ join{\isaliteral{5F}{\isacharunderscore}}commute\ join{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ meet{\isaliteral{5F}{\isacharunderscore}}connection\ meet{\isaliteral{5F}{\isacharunderscore}}related{\isadigit{2}}\ total{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}l\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}r{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse% -\ \isacommand{note}\isamarkupfalse% -\ total\isanewline -\ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{by}\isamarkupfalse% -\ blast\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -The locale hierarchy is now as shown in - Figure~\ref{fig:lattices}(c).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locale interpretation is \emph{dynamic}. The statement - \isakeyword{sublocale} $l_1 \subseteq l_2$ will not just add the - current conclusions of $l_2$ to $l_1$. Rather the dependency is - stored, and conclusions that will be - added to $l_2$ in future are automatically propagated to $l_1$. - The sublocale relation is transitive --- that is, propagation takes - effect along chains of sublocales. Even cycles in the sublocale relation are - supported, as long as these cycles do not lead to infinite chains. - Details are discussed in the technical report \cite{Ballarin2006a}. - See also Section~\ref{sec:infinite-chains} of this tutorial.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/Examples1.tex --- a/doc-src/Locales/Locales/document/Examples1.tex Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,159 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Examples{\isadigit{1}}}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Examples{\isadigit{1}}\isanewline -\isakeyword{imports}\ Examples\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\vspace{-5ex}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Use of Locales in Theories and Proofs - \label{sec:interpretation}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locales can be interpreted in the contexts of theories and - structured proofs. These interpretations are dynamic, too. - Conclusions of locales will be propagated to the current theory or - the current proof context.% -\footnote{Strictly speaking, only interpretation in theories is - dynamic since it is not possible to change locales or the locale - hierarchy from within a proof.} - The focus of this section is on - interpretation in theories, but we will also encounter - interpretations in proofs, in - Section~\ref{sec:local-interpretation}. - - As an example, consider the type of integers \isa{int}. The - relation \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} is a total order over \isa{int}. We start - with the interpretation that \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} is a partial order. The - facilities of the interpretation command are explored gradually in - three versions.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{First Version: Replacement of Parameters Only - \label{sec:po-first}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The command \isakeyword{interpretation} is for the interpretation of - locale in theories. In the following example, the parameter of locale - \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} is replaced by \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} and the locale instance is interpreted in the current - theory.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{interpretation}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptxt}% -\normalsize - The argument of the command is a simple \emph{locale expression} - consisting of the name of the interpreted locale, which is - preceded by the qualifier \isa{int{\isaliteral{3A}{\isacharcolon}}} and succeeded by a - white-space-separated list of terms, which provide a full - instantiation of the locale parameters. The parameters are referred - to by order of declaration, which is also the order in which - \isakeyword{print\_locale} outputs them. The locale has only a - single parameter, hence the list of instantiation terms is a - singleton. - - The command creates the goal - \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}% -\end{isabelle} which can be shown easily:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\ auto% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -The effect of the command is that instances of all - conclusions of the locale are available in the theory, where names - are prefixed by the qualifier. For example, transitivity for \isa{int} is named \isa{int{\isaliteral{2E}{\isachardot}}trans} and is the following - theorem: - \begin{isabelle}% -\ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}z% -\end{isabelle} - It is not possible to reference this theorem simply as \isa{trans}. This prevents unwanted hiding of existing theorems of the - theory by an interpretation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Second Version: Replacement of Definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Not only does the above interpretation qualify theorem names. - The prefix \isa{int} is applied to all names introduced in locale - conclusions including names introduced in definitions. The - qualified name \isa{int{\isaliteral{2E}{\isachardot}}less} is short for - the interpretation of the definition, which is \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}}. - Qualified name and expanded form may be used almost - interchangeably.% -\footnote{Since \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} is polymorphic, for \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}} a - more general type will be inferred than for \isa{int{\isaliteral{2E}{\isachardot}}less} which - is over type \isa{int}.} - The latter is preferred on output, as for example in the theorem - \isa{int{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans}: \begin{isabelle}% -\ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline -\isaindent{\ \ }{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}z% -\end{isabelle} - Both notations for the strict order are not satisfactory. The - constant \isa{op\ {\isaliteral{3C}{\isacharless}}} is the strict order for \isa{int}. - In order to allow for the desired replacement, interpretation - accepts \emph{equations} in addition to the parameter instantiation. - These follow the locale expression and are indicated with the - keyword \isakeyword{where}. This is the revised interpretation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/Examples2.tex --- a/doc-src/Locales/Locales/document/Examples2.tex Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,93 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Examples{\isadigit{2}}}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Examples{\isadigit{2}}\isanewline -\isakeyword{imports}\ Examples\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\vspace{-5ex}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{interpretation}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5B}{\isacharbrackleft}}int{\isaliteral{2C}{\isacharcomma}}\ int{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}less\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3C}{\isacharless}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}% -\begin{isamarkuptxt}% -\normalsize The goals are now: - \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\isanewline -\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3C}{\isacharless}}\ y{\isaliteral{29}{\isacharparenright}}% -\end{isabelle} - The proof that~\isa{{\isaliteral{5C3C6C653E}{\isasymle}}} is a partial order is as above.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\ auto% -\begin{isamarkuptxt}% -\normalsize The second goal is shown by unfolding the - definition of \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less}.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3C}{\isacharless}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}def\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{60}{\isacharbackquoteopen}}partial{\isaliteral{5F}{\isacharunderscore}}order\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ auto\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -Note that the above proof is not in the context of the - interpreted locale. Hence, the premise of \isa{partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}def} is discharged manually with \isa{OF}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/Examples3.tex --- a/doc-src/Locales/Locales/document/Examples3.tex Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,975 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Examples{\isadigit{3}}}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Examples{\isadigit{3}}\isanewline -\isakeyword{imports}\ Examples\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\vspace{-5ex}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Third Version: Local Interpretation - \label{sec:local-interpretation}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In the above example, the fact that \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} is a partial - order for the integers was used in the second goal to - discharge the premise in the definition of \isa{op\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}}. In - general, proofs of the equations not only may involve definitions - from the interpreted locale but arbitrarily complex arguments in the - context of the locale. Therefore it would be convenient to have the - interpreted locale conclusions temporarily available in the proof. - This can be achieved by a locale interpretation in the proof body. - The command for local interpretations is \isakeyword{interpret}. We - repeat the example from the previous section to illustrate this.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{interpretation}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}less\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3C}{\isacharless}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\ auto\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{interpret}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5B}{\isacharbrackleft}}int{\isaliteral{2C}{\isacharcomma}}\ int{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}less\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3C}{\isacharless}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ int{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}def\ \isacommand{by}\isamarkupfalse% -\ auto\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -The inner interpretation is immediate from the preceding fact - and proved by assumption (Isar short hand ``.''). It enriches the - local proof context by the theorems - also obtained in the interpretation from Section~\ref{sec:po-first}, - and \isa{int{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}def} may directly be used to unfold the - definition. Theorems from the local interpretation disappear after - leaving the proof context --- that is, after the succeeding - \isakeyword{next} or \isakeyword{qed} statement.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Further Interpretations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Further interpretations are necessary for - the other locales. In \isa{lattice} the operations~\isa{{\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}} - and~\isa{{\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}} are substituted by \isa{min} - and \isa{max}. The entire proof for the - interpretation is reproduced to give an example of a more - elaborate interpretation proof. Note that the equations are named - so they can be used in a later example.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{interpretation}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ lattice\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isakeyword{where}\ int{\isaliteral{5F}{\isacharunderscore}}min{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}meet\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ min\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ int{\isaliteral{5F}{\isacharunderscore}}max{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}join\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ max\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}lattice\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptxt}% -\normalsize We have already shown that this is a partial - order,% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales% -\begin{isamarkuptxt}% -\normalsize hence only the lattice axioms remain to be - shown. - \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}inf{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}inf\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ x\ y\ inf\isanewline -\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}sup{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}sup\ op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ x\ y\ sup% -\end{isabelle} - By \isa{is{\isaliteral{5F}{\isacharunderscore}}inf} and \isa{is{\isaliteral{5F}{\isacharunderscore}}sup},% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold\ int{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}def\ int{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}% -\begin{isamarkuptxt}% -\normalsize the goals are transformed to these - statements: - \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}inf{\isaliteral{5C3C6C653E}{\isasymle}}x{\isaliteral{2E}{\isachardot}}\ inf\ {\isaliteral{5C3C6C653E}{\isasymle}}\ y\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ z\ {\isaliteral{5C3C6C653E}{\isasymle}}\ x\ {\isaliteral{5C3C616E643E}{\isasymand}}\ z\ {\isaliteral{5C3C6C653E}{\isasymle}}\ y\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ z\ {\isaliteral{5C3C6C653E}{\isasymle}}\ inf{\isaliteral{29}{\isacharparenright}}\isanewline -\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}sup{\isaliteral{5C3C67653E}{\isasymge}}x{\isaliteral{2E}{\isachardot}}\ y\ {\isaliteral{5C3C6C653E}{\isasymle}}\ sup\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C6C653E}{\isasymle}}\ z\ {\isaliteral{5C3C616E643E}{\isasymand}}\ y\ {\isaliteral{5C3C6C653E}{\isasymle}}\ z\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ sup\ {\isaliteral{5C3C6C653E}{\isasymle}}\ z{\isaliteral{29}{\isacharparenright}}% -\end{isabelle} - This is Presburger arithmetic, which can be solved by the - method \isa{arith}.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ arith{\isaliteral{2B}{\isacharplus}}% -\begin{isamarkuptxt}% -\normalsize In order to show the equations, we put ourselves - in a situation where the lattice theorems can be used in a - convenient way.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{interpret}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ lattice\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}meet\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ min\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}bestsimp\ simp{\isaliteral{3A}{\isacharcolon}}\ int{\isaliteral{2E}{\isachardot}}meet{\isaliteral{5F}{\isacharunderscore}}def\ int{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}inf{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}int{\isaliteral{2E}{\isachardot}}join\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ max\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}bestsimp\ simp{\isaliteral{3A}{\isacharcolon}}\ int{\isaliteral{2E}{\isachardot}}join{\isaliteral{5F}{\isacharunderscore}}def\ int{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}sup{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -Next follows that \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} is a total order, again for - the integers.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{interpretation}\isamarkupfalse% -\ int{\isaliteral{3A}{\isacharcolon}}\ total{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\ arith% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -Theorems that are available in the theory at this point are shown in - Table~\ref{tab:int-lattice}. Two points are worth noting: - -\begin{table} -\hrule -\vspace{2ex} -\begin{center} -\begin{tabular}{l} - \isa{int{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}def} from locale \isa{partial{\isaliteral{5F}{\isacharunderscore}}order}: \\ - \quad \isa{{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3C}{\isacharless}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{29}{\isacharparenright}}} \\ - \isa{int{\isaliteral{2E}{\isachardot}}meet{\isaliteral{5F}{\isacharunderscore}}left} from locale \isa{lattice}: \\ - \quad \isa{min\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{3F}{\isacharquery}}x} \\ - \isa{int{\isaliteral{2E}{\isachardot}}join{\isaliteral{5F}{\isacharunderscore}}distr} from locale \isa{distrib{\isaliteral{5F}{\isacharunderscore}}lattice}: \\ - \quad \isa{max\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{28}{\isacharparenleft}}min\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ min\ {\isaliteral{28}{\isacharparenleft}}max\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}max\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{29}{\isacharparenright}}} \\ - \isa{int{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}total} from locale \isa{total{\isaliteral{5F}{\isacharunderscore}}order}: \\ - \quad \isa{{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3C}{\isacharless}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{3C}{\isacharless}}\ {\isaliteral{3F}{\isacharquery}}x} -\end{tabular} -\end{center} -\hrule -\caption{Interpreted theorems for~\isa{{\isaliteral{5C3C6C653E}{\isasymle}}} on the integers.} -\label{tab:int-lattice} -\end{table} - -\begin{itemize} -\item - Locale \isa{distrib{\isaliteral{5F}{\isacharunderscore}}lattice} was also interpreted. Since the - locale hierarchy reflects that total orders are distributive - lattices, the interpretation of the latter was inserted - automatically with the interpretation of the former. In general, - interpretation traverses the locale hierarchy upwards and interprets - all encountered locales, regardless whether imported or proved via - the \isakeyword{sublocale} command. Existing interpretations are - skipped avoiding duplicate work. -\item - The predicate \isa{op\ {\isaliteral{3C}{\isacharless}}} appears in theorem \isa{int{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}total} - although an equation for the replacement of \isa{op\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}} was only - given in the interpretation of \isa{partial{\isaliteral{5F}{\isacharunderscore}}order}. The - interpretation equations are pushed downwards the hierarchy for - related interpretations --- that is, for interpretations that share - the instances of parameters they have in common. -\end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The interpretations for a locale $n$ within the current - theory may be inspected with \isakeyword{print\_interps}~$n$. This - prints the list of instances of $n$, for which interpretations exist. - For example, \isakeyword{print\_interps} \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} - outputs the following: -\begin{small} -\begin{alltt} - int! : partial_order "op \(\le\)" -\end{alltt} -\end{small} - Of course, there is only one interpretation. - The interpretation qualifier on the left is decorated with an - exclamation point. This means that it is mandatory. Qualifiers - can either be \emph{mandatory} or \emph{optional}, designated by - ``!'' or ``?'' respectively. Mandatory qualifiers must occur in a - name reference while optional ones need not. Mandatory qualifiers - prevent accidental hiding of names, while optional qualifiers can be - more convenient to use. For \isakeyword{interpretation}, the - default is ``!''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Locale Expressions \label{sec:expressions}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A map~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} between partial orders~\isa{{\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}} and~\isa{{\isaliteral{5C3C7072656365713E}{\isasympreceq}}} - is called order preserving if \isa{x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y} implies \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}\ x\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ y}. This situation is more complex than those encountered so - far: it involves two partial orders, and it is desirable to use the - existing locale for both. - - A locale for order preserving maps requires three parameters: \isa{le}~(\isakeyword{infixl}~\isa{{\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}}) and \isa{le{\isaliteral{27}{\isacharprime}}}~(\isakeyword{infixl}~\isa{{\isaliteral{5C3C7072656365713E}{\isasympreceq}}}) for the orders and~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} - for the map. - - In order to reuse the existing locale for partial orders, which has - the single parameter~\isa{le}, it must be imported twice, once - mapping its parameter to~\isa{le} from the new locale and once - to~\isa{le{\isaliteral{27}{\isacharprime}}}. This can be achieved with a compound locale - expression. - - In general, a locale expression is a sequence of \emph{locale instances} - separated by~``$\textbf{+}$'' and followed by a \isakeyword{for} - clause. - An instance has the following format: -\begin{quote} - \textit{qualifier} \textbf{:} \textit{locale-name} - \textit{parameter-instantiation} -\end{quote} - We have already seen locale instances as arguments to - \isakeyword{interpretation} in Section~\ref{sec:interpretation}. - As before, the qualifier serves to disambiguate names from - different instances of the same locale. While in - \isakeyword{interpretation} qualifiers default to mandatory, in - import and in the \isakeyword{sublocale} command, they default to - optional. - - Since the parameters~\isa{le} and~\isa{le{\isaliteral{27}{\isacharprime}}} are to be partial - orders, our locale for order preserving maps will import the these - instances: -\begin{small} -\begin{alltt} - le: partial_order le - le': partial_order le' -\end{alltt} -\end{small} - For matter of convenience we choose to name parameter names and - qualifiers alike. This is an arbitrary decision. Technically, qualifiers - and parameters are unrelated. - - Having determined the instances, let us turn to the \isakeyword{for} - clause. It serves to declare locale parameters in the same way as - the context element \isakeyword{fixes} does. Context elements can - only occur after the import section, and therefore the parameters - referred to in the instances must be declared in the \isakeyword{for} - clause. The \isakeyword{for} clause is also where the syntax of these - parameters is declared. - - Two context elements for the map parameter~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} and the - assumptions that it is order preserving complete the locale - declaration.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ order{\isaliteral{5F}{\isacharunderscore}}preserving\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ le{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ le\ {\isaliteral{2B}{\isacharplus}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ le{\isaliteral{27}{\isacharprime}}\isanewline -\ \ \ \ \ \ \isakeyword{for}\ le\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{and}\ le{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7072656365713E}{\isasympreceq}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\isanewline -\ \ \ \ \isakeyword{fixes}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isanewline -\ \ \ \ \isakeyword{assumes}\ hom{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ x\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ y{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -Here are examples of theorems that are - available in the locale: - - \hspace*{1em}\isa{hom{\isaliteral{5F}{\isacharunderscore}}le}: \isa{{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{3F}{\isacharquery}}y} - - \hspace*{1em}\isa{le{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans}: \isa{{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C73717375627365743E}{\isasymsqsubset}}\ {\isaliteral{3F}{\isacharquery}}z} - - \hspace*{1em}\isa{le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans}: - \begin{isabelle}% -\ \ \ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline -\isaindent{\ \ \ \ }{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ op\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{3F}{\isacharquery}}z% -\end{isabelle} - While there is infix syntax for the strict operation associated to - \isa{op\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}}, there is none for the strict version of \isa{op\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}}. The abbreviation \isa{less} with its infix syntax is only - available for the original instance it was declared for. We may - introduce the abbreviation \isa{less{\isaliteral{27}{\isacharprime}}} with infix syntax~\isa{{\isaliteral{5C3C707265633E}{\isasymprec}}} - with the following declaration:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{abbreviation}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ order{\isaliteral{5F}{\isacharunderscore}}preserving{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ less{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C707265633E}{\isasymprec}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}less{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}less\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -Now the theorem is displayed nicely as - \isa{le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans}: - \begin{isabelle}% -\ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C707265633E}{\isasymprec}}\ {\isaliteral{3F}{\isacharquery}}y{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{3F}{\isacharquery}}z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C707265633E}{\isasymprec}}\ {\isaliteral{3F}{\isacharquery}}z% -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -There are short notations for locale expressions. These are - discussed in the following.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Default Instantiations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -It is possible to omit parameter instantiations. The - instantiation then defaults to the name of - the parameter itself. For example, the locale expression \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} is short for \isa{partial{\isaliteral{5F}{\isacharunderscore}}order\ le}, since the - locale's single parameter is~\isa{le}. We took advantage of this - in the \isakeyword{sublocale} declarations of - Section~\ref{sec:changing-the-hierarchy}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Implicit Parameters \label{sec:implicit-parameters}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In a locale expression that occurs within a locale - declaration, omitted parameters additionally extend the (possibly - empty) \isakeyword{for} clause. - - The \isakeyword{for} clause is a general construct of Isabelle/Isar - to mark names occurring in the preceding declaration as ``arbitrary - but fixed''. This is necessary for example, if the name is already - bound in a surrounding context. In a locale expression, names - occurring in parameter instantiations should be bound by a - \isakeyword{for} clause whenever these names are not introduced - elsewhere in the context --- for example, on the left hand side of a - \isakeyword{sublocale} declaration. - - There is an exception to this rule in locale declarations, where the - \isakeyword{for} clause serves to declare locale parameters. Here, - locale parameters for which no parameter instantiation is given are - implicitly added, with their mixfix syntax, at the beginning of the - \isakeyword{for} clause. For example, in a locale declaration, the - expression \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} is short for -\begin{small} -\begin{alltt} - partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.} -\end{alltt} -\end{small} - This short hand was used in the locale declarations throughout - Section~\ref{sec:import}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following locale declarations provide more examples. A - map~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} is a lattice homomorphism if it preserves meet and - join.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ lattice{\isaliteral{5F}{\isacharunderscore}}hom\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ le{\isaliteral{3A}{\isacharcolon}}\ lattice\ {\isaliteral{2B}{\isacharplus}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{3A}{\isacharcolon}}\ lattice\ le{\isaliteral{27}{\isacharprime}}\ \isakeyword{for}\ le{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7072656365713E}{\isasympreceq}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\isanewline -\ \ \ \ \isakeyword{fixes}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isanewline -\ \ \ \ \isakeyword{assumes}\ hom{\isaliteral{5F}{\isacharunderscore}}meet{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}meet\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ hom{\isaliteral{5F}{\isacharunderscore}}join{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}join\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -The parameter instantiation in the first instance of \isa{lattice} is omitted. This causes the parameter~\isa{le} to be - added to the \isakeyword{for} clause, and the locale has - parameters~\isa{le},~\isa{le{\isaliteral{27}{\isacharprime}}} and, of course,~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}}. - - Before turning to the second example, we complete the locale by - providing infix syntax for the meet and join operations of the - second lattice.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{context}\isamarkupfalse% -\ lattice{\isaliteral{5F}{\isacharunderscore}}hom\isanewline -\ \ \isakeyword{begin}\isanewline -\ \ \isacommand{abbreviation}\isamarkupfalse% -\ meet{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}{\isaliteral{27}{\isacharprime}}{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}meet{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}meet{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{abbreviation}\isamarkupfalse% -\ join{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}{\isaliteral{27}{\isacharprime}}{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}join{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}join{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -The next example makes radical use of the short hand - facilities. A homomorphism is an endomorphism if both orders - coincide.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ lattice{\isaliteral{5F}{\isacharunderscore}}end\ {\isaliteral{3D}{\isacharequal}}\ lattice{\isaliteral{5F}{\isacharunderscore}}hom\ {\isaliteral{5F}{\isacharunderscore}}\ le% -\begin{isamarkuptext}% -The notation~\isa{{\isaliteral{5F}{\isacharunderscore}}} enables to omit a parameter in a - positional instantiation. The omitted parameter,~\isa{le} becomes - the parameter of the declared locale and is, in the following - position, used to instantiate the second parameter of \isa{lattice{\isaliteral{5F}{\isacharunderscore}}hom}. The effect is that of identifying the first in second - parameter of the homomorphism locale.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The inheritance diagram of the situation we have now is shown - in Figure~\ref{fig:hom}, where the dashed line depicts an - interpretation which is introduced below. Parameter instantiations - are indicated by $\sqsubseteq \mapsto \preceq$ etc. By looking at - the inheritance diagram it would seem - that two identical copies of each of the locales \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} and \isa{lattice} are imported by \isa{lattice{\isaliteral{5F}{\isacharunderscore}}end}. This is not the case! Inheritance paths with - identical morphisms are automatically detected and - the conclusions of the respective locales appear only once. - -\begin{figure} -\hrule \vspace{2ex} -\begin{center} -\begin{tikzpicture} - \node (o) at (0,0) {\isa{partial{\isaliteral{5F}{\isacharunderscore}}order}}; - \node (oh) at (1.5,-2) {\isa{order{\isaliteral{5F}{\isacharunderscore}}preserving}}; - \node (oh1) at (1.5,-0.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; - \node (oh2) at (0,-1.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; - \node (l) at (-1.5,-2) {\isa{lattice}}; - \node (lh) at (0,-4) {\isa{lattice{\isaliteral{5F}{\isacharunderscore}}hom}}; - \node (lh1) at (0,-2.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; - \node (lh2) at (-1.5,-3.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; - \node (le) at (0,-6) {\isa{lattice{\isaliteral{5F}{\isacharunderscore}}end}}; - \node (le1) at (0,-4.8) - [anchor=west]{$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; - \node (le2) at (0,-5.2) - [anchor=west]{$\scriptscriptstyle \preceq \mapsto \sqsubseteq$}; - \draw (o) -- (l); - \draw[dashed] (oh) -- (lh); - \draw (lh) -- (le); - \draw (o) .. controls (oh1.south west) .. (oh); - \draw (o) .. controls (oh2.north east) .. (oh); - \draw (l) .. controls (lh1.south west) .. (lh); - \draw (l) .. controls (lh2.north east) .. (lh); -\end{tikzpicture} -\end{center} -\hrule -\caption{Hierarchy of Homomorphism Locales.} -\label{fig:hom} -\end{figure}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -It can be shown easily that a lattice homomorphism is order - preserving. As the final example of this section, a locale - interpretation is used to assert this:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{sublocale}\isamarkupfalse% -\ lattice{\isaliteral{5F}{\isacharunderscore}}hom\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ order{\isaliteral{5F}{\isacharunderscore}}preserving\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ le{\isaliteral{2E}{\isachardot}}join{\isaliteral{5F}{\isacharunderscore}}connection{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ y\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ hom{\isaliteral{5F}{\isacharunderscore}}join\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ x\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ le{\isaliteral{27}{\isacharprime}}{\isaliteral{2E}{\isachardot}}join{\isaliteral{5F}{\isacharunderscore}}connection{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Theorems and other declarations --- syntax, in particular --- from - the locale \isa{order{\isaliteral{5F}{\isacharunderscore}}preserving} are now active in \isa{lattice{\isaliteral{5F}{\isacharunderscore}}hom}, for example - \isa{hom{\isaliteral{5F}{\isacharunderscore}}le}: - \begin{isabelle}% -\ \ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ {\isaliteral{3F}{\isacharquery}}y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C7072656365713E}{\isasympreceq}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{3F}{\isacharquery}}y% -\end{isabelle} - This theorem will be useful in the following section.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Conditional Interpretation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -There are situations where an interpretation is not possible - in the general case since the desired property is only valid if - certain conditions are fulfilled. Take, for example, the function - \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}i{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ i} that scales its argument by a constant factor. - This function is order preserving (and even a lattice endomorphism) - with respect to \isa{op\ {\isaliteral{5C3C6C653E}{\isasymle}}} provided \isa{n\ {\isaliteral{5C3C67653E}{\isasymge}}\ {\isadigit{0}}}. - - It is not possible to express this using a global interpretation, - because it is in general unspecified whether~\isa{n} is - non-negative, but one may make an interpretation in an inner context - of a proof where full information is available. - This is not fully satisfactory either, since potentially - interpretations may be required to make interpretations in many - contexts. What is - required is an interpretation that depends on the condition --- and - this can be done with the \isakeyword{sublocale} command. For this - purpose, we introduce a locale for the condition.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ non{\isaliteral{5F}{\isacharunderscore}}negative\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ \isakeyword{fixes}\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ int\isanewline -\ \ \ \ \isakeyword{assumes}\ non{\isaliteral{5F}{\isacharunderscore}}neg{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{0}}\ {\isaliteral{5C3C6C653E}{\isasymle}}\ n{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -It is again convenient to make the interpretation in an - incremental fashion, first for order preserving maps, the for - lattice endomorphisms.% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{sublocale}\isamarkupfalse% -\ non{\isaliteral{5F}{\isacharunderscore}}negative\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\isanewline -\ \ \ \ \ \ order{\isaliteral{5F}{\isacharunderscore}}preserving\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}i{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{using}\isamarkupfalse% -\ non{\isaliteral{5F}{\isacharunderscore}}neg\ \isacommand{by}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\ {\isaliteral{28}{\isacharparenleft}}rule\ mult{\isaliteral{5F}{\isacharunderscore}}left{\isaliteral{5F}{\isacharunderscore}}mono{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -While the proof of the previous interpretation - is straightforward from monotonicity lemmas for~\isa{op\ {\isaliteral{2A}{\isacharasterisk}}}, the - second proof follows a useful pattern.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{sublocale}\isamarkupfalse% -\ non{\isaliteral{5F}{\isacharunderscore}}negative\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ lattice{\isaliteral{5F}{\isacharunderscore}}end\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C6C653E}{\isasymle}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}i{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}unfold{\isaliteral{5F}{\isacharunderscore}}locales{\isaliteral{2C}{\isacharcomma}}\ unfold\ int{\isaliteral{5F}{\isacharunderscore}}min{\isaliteral{5F}{\isacharunderscore}}eq\ int{\isaliteral{5F}{\isacharunderscore}}max{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{29}{\isacharparenright}}% -\begin{isamarkuptxt}% -\normalsize Unfolding the locale predicates \emph{and} the - interpretation equations immediately yields two subgoals that - reflect the core conjecture. - \begin{isabelle}% -\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ min\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ min\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2A}{\isacharasterisk}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2A}{\isacharasterisk}}\ y{\isaliteral{29}{\isacharparenright}}\isanewline -\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ max\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ max\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2A}{\isacharasterisk}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2A}{\isacharasterisk}}\ y{\isaliteral{29}{\isacharparenright}}% -\end{isabelle} - It is now necessary to show, in the context of \isa{non{\isaliteral{5F}{\isacharunderscore}}negative}, that multiplication by~\isa{n} commutes with - \isa{min} and \isa{max}.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{qed}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}auto\ simp{\isaliteral{3A}{\isacharcolon}}\ hom{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{29}{\isacharparenright}}% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -The lemma \isa{hom{\isaliteral{5F}{\isacharunderscore}}le} - simplifies a proof that would have otherwise been lengthy and we may - consider making it a default rule for the simplifier:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemmas}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ order{\isaliteral{5F}{\isacharunderscore}}preserving{\isaliteral{29}{\isacharparenright}}\ hom{\isaliteral{5F}{\isacharunderscore}}le\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}% -\isamarkupsubsection{Avoiding Infinite Chains of Interpretations - \label{sec:infinite-chains}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Similar situations arise frequently in formalisations of - abstract algebra where it is desirable to express that certain - constructions preserve certain properties. For example, polynomials - over rings are rings, or --- an example from the domain where the - illustrations of this tutorial are taken from --- a partial order - may be obtained for a function space by point-wise lifting of the - partial order of the co-domain. This corresponds to the following - interpretation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -\ \ % -\endisadelimvisible -% -\isatagvisible -\isacommand{sublocale}\isamarkupfalse% -\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ f{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{oops}\isamarkupfalse% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -Unfortunately this is a cyclic interpretation that leads to an - infinite chain, namely - \begin{isabelle}% -\ \ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\isanewline -\isaindent{\ \ }\ \ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x\ y{\isaliteral{2E}{\isachardot}}\ f\ x\ y\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}% -\end{isabelle} - and the interpretation is rejected. - - Instead it is necessary to declare a locale that is logically - equivalent to \isa{partial{\isaliteral{5F}{\isacharunderscore}}order} but serves to collect facts - about functions spaces where the co-domain is a partial order, and - to make the interpretation in its context:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ fun{\isaliteral{5F}{\isacharunderscore}}partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{3D}{\isacharequal}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\isanewline -\isanewline -\ \ \isacommand{sublocale}\isamarkupfalse% -\ fun{\isaliteral{5F}{\isacharunderscore}}partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\isanewline -\ \ \ \ \ \ f{\isaliteral{3A}{\isacharcolon}}\ partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\ {\isaliteral{28}{\isacharparenleft}}fast{\isaliteral{2C}{\isacharcomma}}rule{\isaliteral{2C}{\isacharcomma}}fast{\isaliteral{2C}{\isacharcomma}}blast\ intro{\isaliteral{3A}{\isacharcolon}}\ trans{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -It is quite common in abstract algebra that such a construction - maps a hierarchy of algebraic structures (or specifications) to a - related hierarchy. By means of the same lifting, a function space - is a lattice if its co-domain is a lattice:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{locale}\isamarkupfalse% -\ fun{\isaliteral{5F}{\isacharunderscore}}lattice\ {\isaliteral{3D}{\isacharequal}}\ fun{\isaliteral{5F}{\isacharunderscore}}partial{\isaliteral{5F}{\isacharunderscore}}order\ {\isaliteral{2B}{\isacharplus}}\ lattice\isanewline -\isanewline -\ \ \isacommand{sublocale}\isamarkupfalse% -\ fun{\isaliteral{5F}{\isacharunderscore}}lattice\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ f{\isaliteral{3A}{\isacharcolon}}\ lattice\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ f\ g\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}inf\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{29}{\isacharparenright}}\ f\ g\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C7371696E7465723E}{\isasymsqinter}}\ g\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}infI{\isaliteral{29}{\isacharparenright}}\ \isacommand{apply}\isamarkupfalse% -\ rule{\isaliteral{2B}{\isacharplus}}\ \isacommand{apply}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}drule\ spec{\isaliteral{2C}{\isacharcomma}}\ assumption{\isaliteral{29}{\isacharparenright}}{\isaliteral{2B}{\isacharplus}}\ \isacommand{done}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}inf{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}inf\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{29}{\isacharparenright}}\ f\ g\ inf{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ fast\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ f\ g\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}sup\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{29}{\isacharparenright}}\ f\ g\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C7371756E696F6E3E}{\isasymsqunion}}\ g\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ is{\isaliteral{5F}{\isacharunderscore}}supI{\isaliteral{29}{\isacharparenright}}\ \isacommand{apply}\isamarkupfalse% -\ rule{\isaliteral{2B}{\isacharplus}}\ \isacommand{apply}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}drule\ spec{\isaliteral{2C}{\isacharcomma}}\ assumption{\isaliteral{29}{\isacharparenright}}{\isaliteral{2B}{\isacharplus}}\ \isacommand{done}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}sup{\isaliteral{2E}{\isachardot}}\ partial{\isaliteral{5F}{\isacharunderscore}}order{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}sup\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}f\ g{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{5C3C737173756273657465713E}{\isasymsqsubseteq}}\ g\ x{\isaliteral{29}{\isacharparenright}}\ f\ g\ sup{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ fast\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{Further Reading% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -More information on locales and their interpretation is - available. For the locale hierarchy of import and interpretation - dependencies see~\cite{Ballarin2006a}; interpretations in theories - and proofs are covered in~\cite{Ballarin2006b}. In the latter, I - show how interpretation in proofs enables to reason about families - of algebraic structures, which cannot be expressed with locales - directly. - - Haftmann and Wenzel~\cite{HaftmannWenzel2007} overcome a restriction - of axiomatic type classes through a combination with locale - interpretation. The result is a Haskell-style class system with a - facility to generate ML and Haskell code. Classes are sufficient for - simple specifications with a single type parameter. The locales for - orders and lattices presented in this tutorial fall into this - category. Order preserving maps, homomorphisms and vector spaces, - on the other hand, do not. - - The locales reimplementation for Isabelle 2009 provides, among other - improvements, a clean integration with Isabelle/Isar's local theory - mechanisms, which are described in another paper by Haftmann and - Wenzel~\cite{HaftmannWenzel2009}. - - The original work of Kamm\"uller on locales~\cite{KammullerEtAl1999} - may be of interest from a historical perspective. My previous - report on locales and locale expressions~\cite{Ballarin2004a} - describes a simpler form of expressions than available now and is - outdated. The mathematical background on orders and lattices is - taken from Jacobson's textbook on algebra~\cite[Chapter~8]{Jacobson1985}. - - The sources of this tutorial, which include all proofs, are - available with the Isabelle distribution at - \url{http://isabelle.in.tum.de}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{table} -\hrule -\vspace{2ex} -\begin{center} -\begin{tabular}{l>$c<$l} - \multicolumn{3}{l}{Miscellaneous} \\ - - \textit{attr-name} & ::= - & \textit{name} $|$ \textit{attribute} $|$ - \textit{name} \textit{attribute} \\ - \textit{qualifier} & ::= - & \textit{name} [``\textbf{?}'' $|$ ``\textbf{!}''] \\[2ex] - - \multicolumn{3}{l}{Context Elements} \\ - - \textit{fixes} & ::= - & \textit{name} [ ``\textbf{::}'' \textit{type} ] - [ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ - \textit{mixfix} ] \\ -\begin{comment} - \textit{constrains} & ::= - & \textit{name} ``\textbf{::}'' \textit{type} \\ -\end{comment} - \textit{assumes} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ -\begin{comment} - \textit{defines} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ - \textit{notes} & ::= - & [ \textit{attr-name} ``\textbf{=}'' ] - ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ -\end{comment} - - \textit{element} & ::= - & \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ -\begin{comment} - & | - & \textbf{constrains} \textit{constrains} - ( \textbf{and} \textit{constrains} )$^*$ \\ -\end{comment} - & | - & \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\[2ex] -%\begin{comment} -% & | -% & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ -% & | -% & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ -%\end{comment} - - \multicolumn{3}{l}{Locale Expressions} \\ - - \textit{pos-insts} & ::= - & ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\ - \textit{named-insts} & ::= - & \textbf{where} \textit{name} ``\textbf{=}'' \textit{term} - ( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\ - \textit{instance} & ::= - & [ \textit{qualifier} ``\textbf{:}'' ] - \textit{name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\ - \textit{expression} & ::= - & \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$ - [ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex] - - \multicolumn{3}{l}{Declaration of Locales} \\ - - \textit{locale} & ::= - & \textit{element}$^+$ \\ - & | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ - \textit{toplevel} & ::= - & \textbf{locale} \textit{name} [ ``\textbf{=}'' - \textit{locale} ] \\[2ex] - - \multicolumn{3}{l}{Interpretation} \\ - - \textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] - \textit{prop} \\ - \textit{equations} & ::= & \textbf{where} \textit{equation} ( \textbf{and} - \textit{equation} )$^*$ \\ - \textit{toplevel} & ::= - & \textbf{sublocale} \textit{name} ( ``$<$'' $|$ - ``$\subseteq$'' ) \textit{expression} \textit{proof} \\ - & | - & \textbf{interpretation} - \textit{expression} [ \textit{equations} ] \textit{proof} \\ - & | - & \textbf{interpret} - \textit{expression} \textit{proof} \\[2ex] - - \multicolumn{3}{l}{Diagnostics} \\ - - \textit{toplevel} & ::= - & \textbf{print\_locales} \\ - & | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{name} \\ - & | & \textbf{print\_interps} \textit{name} -\end{tabular} -\end{center} -\hrule -\caption{Syntax of Locale Commands.} -\label{tab:commands} -\end{table}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\textbf{Revision History.} For the present third revision of - the tutorial, much of the explanatory text - was rewritten. Inheritance of interpretation equations is - available with the forthcoming release of Isabelle, which at the - time of editing these notes is expected for the end of 2009. - The second revision accommodates changes introduced by the locales - reimplementation for Isabelle 2009. Most notably locale expressions - have been generalised from renaming to instantiation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, - Randy Pollack, Andreas Schropp, Christian Sternagel and Makarius Wenzel - have made - useful comments on earlier versions of this document. The section - on conditional interpretation was inspired by a number of e-mail - enquiries the author received from locale users, and which suggested - that this use case is important enough to deserve explicit - explanation. The term \emph{conditional interpretation} is due to - Larry Paulson.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/GCD.tex diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/root.bib --- a/doc-src/Locales/Locales/document/root.bib Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,117 +0,0 @@ -@unpublished{IsarRef, - author = "Markus Wenzel", - title = "The {Isabelle/Isar} Reference Manual", - note = "Part of the Isabelle distribution, \url{http://isabelle.in.tum.de/doc/isar-ref.pdf}." -} - -@book {Jacobson1985, - author = "Nathan Jacobson", - title = "Basic Algebra", - volume = "I", - publisher = "Freeman", - edition = "2nd", - year = 1985, - available = { CB } -} - -% TYPES 2006 - -@inproceedings{HaftmannWenzel2007, - author = "Florian Haftmann and Makarius Wenzel", - title = "Constructive Type Classes in {Isabelle}", - pages = "160--174", - crossref = "AltenkirchMcBride2007", - available = { CB } -} - -@proceedings{AltenkirchMcBride2007, - editor = "Thorsten Altenkirch and Connor McBride", - title = "Types for Proofs and Programs, TYPES 2006, Nottingham, UK", - booktitle = "Types for Proofs and Programs, TYPES 2006, Nottingham, UK", - publisher = "Springer", - series = "LNCS 4502", - year = 2007 -} - - -@techreport{Ballarin2006a, - author = "Clemens Ballarin", - title = "Interpretation of Locales in {Isabelle}: Managing Dependencies between Locales", - institution = "Technische Universit{\"a}t M{\"u}nchen", - number = "TUM-I0607", - year = 2006 -} - -% TYPES 2003 - -@inproceedings{Ballarin2004a, - author = "Clemens Ballarin", - title = "Locales and Locale Expressions in {Isabelle/Isar}", - pages = "34--50", - crossref = "BerardiEtAl2004" -} - -@proceedings{BerardiEtAl2004, - editor = "Stefano Berardi and Mario Coppo and Ferruccio Damiani", - title = "Types for Proofs and Programs, TYPES 2003, Torino, Italy", - booktitle = "Types for Proofs and Programs, TYPES 2003, Torino, Italy", - publisher = "Springer", - series = "LNCS 3085", - year = 2004 -} - -% TYPES 2008 - -@inproceedings{HaftmannWenzel2009, - author = "Florian Haftmann and Makarius Wenzel", - title = "Local theory specifications in {Isabelle}/{Isar}", - pages = "153--168", - crossref = "BerardiEtAl2009" -} - -@proceedings{BerardiEtAl2009, - editor = "Stefano Berardi and Ferruccio Damiani and Ugo de Liguoro", - title = "Types for Proofs and Programs, TYPES 2008, Torino, Italy", - booktitle = "Types for Proofs and Programs, TYPES 2008, Torino, Italy", - series = "LNCS 5497", - publisher = "Springer", - year = 2009 -} - -% MKM 2006 - -@inproceedings{Ballarin2006b, - author = "Clemens Ballarin", - title = "Interpretation of Locales in {Isabelle}: Theories and Proof Contexts", - pages = "31--43", - crossref = "BorweinFarmer2006" -} - -@proceedings{BorweinFarmer2006, - editor = "Jonathan M. Borwein and William M. Farmer", - title = "Mathematical knowledge management, MKM 2006, Wokingham, UK", - booktitle = "Mathematical knowledge management, MKM 2006, Wokingham, UK", - series = "LNCS 4108", - publisher = "Springer", - year = 2006, - available = { CB } -} - -% TPHOLs 1999 - -@inproceedings{KammullerEtAl1999, - author = "Florian Kamm{\"u}ller and Markus Wenzel and Lawrence C. Paulson", - title = "Locales: A Sectioning Concept for {Isabelle}", - pages = "149--165", - crossref = "BertotEtAl1999", - available = { CB } -} - -@book{BertotEtAl1999, - editor = "Y. Bertot and G. Dowek and A. Hirschowitz and C. Paulin and L. Th{\'e}ry", - title = "Theorem Proving in Higher Order Logics: TPHOLs'99, Nice, France", - booktitle = "Theorem Proving in Higher Order Logics: TPHOLs'99, Nice, France", - publisher = "Springer", - series = "LNCS 1690", - year = 1999 -} diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/root.tex --- a/doc-src/Locales/Locales/document/root.tex Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,62 +0,0 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{amsmath} -\usepackage{../../../../lib/texinputs/isabelle,../../../../lib/texinputs/isabellesym} -\usepackage{verbatim} -\usepackage{alltt} -\usepackage{array} - -\usepackage{amssymb} - -\usepackage{../../../pdfsetup} - -\usepackage{ifpdf} -\ifpdf\relax\else\def\pgfsysdriver{pgfsys-dvi.def}\fi -\usepackage{tikz} -\usepackage{subfigure} - -\isadroptag{theory} -\isafoldtag{proof} - -% urls in roman style, theory text in typewriter -\urlstyle{rm} -\isabellestyle{tt} - - -\begin{document} - -\title{Tutorial to Locales and Locale Interpretation% -\thanks{Published in L.~Lamb\'an, A.~Romero, J.~Rubio, editors, {\em Contribuciones Cient\'{\i}ficas en honor de Mirian Andr\'es.} Servicio de Publicaciones de la Universidad de La Rioja, Logro\~no, Spain, 2010. Reproduced by permission.}} -\author{Clemens Ballarin} -\date{} - -\maketitle - -\begin{abstract} - Locales are Isabelle's approach for dealing with parametric - theories. They have been designed as a module system for a - theorem prover that can adequately represent the complex - inter-dependencies between structures found in abstract algebra, but - have proven fruitful also in other applications --- for example, - software verification. - - Both design and implementation of locales have evolved considerably - since Kamm\"uller did his initial experiments. Today, locales - are a simple yet powerful extension of the Isar proof language. - The present tutorial covers all major facilities of locales. It is - intended for locale novices; familiarity with Isabelle and Isar is - presumed. -\end{abstract} - -\parindent 0pt\parskip 0.5ex - -\input{session} - -\bibliographystyle{abbrv} -\bibliography{root} - -\end{document} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Locales/document/session.tex --- a/doc-src/Locales/Locales/document/session.tex Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,12 +0,0 @@ -\input{Examples.tex} - -\input{Examples1.tex} - -\input{Examples2.tex} - -\input{Examples3.tex} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/Makefile --- a/doc-src/Locales/Makefile Mon Aug 27 21:04:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,39 +0,0 @@ - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = locales - -FILES = Locales/document/root.tex Locales/document/root.bib \ - Locales/document/session.tex Locales/document/Examples.tex \ - Locales/document/Examples1.tex Locales/document/Examples2.tex \ - Locales/document/Examples3.tex \ - ../../lib/texinputs/isabelle.sty ../../lib/texinputs/isabellesym.sty ../pdfsetup.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) - cd Locales/document && \ - $(LATEX) root && \ - $(BIBTEX) root && \ - $(LATEX) root && \ - $(LATEX) root && \ - $(LATEX) root - mv Locales/document/root.dvi $(NAME).dvi - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) - cd Locales/document && \ - $(PDFLATEX) root && \ - $(BIBTEX) root && \ - $(PDFLATEX) root && \ - $(PDFLATEX) root && \ - $(PDFLATEX) root - mv Locales/document/root.pdf $(NAME).pdf diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/document/build --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/document/build Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,16 @@ +#!/bin/bash + +set -e + +FORMAT="$1" +VARIANT="$2" + +"$ISABELLE_TOOL" latex -o sty +"$ISABELLE_TOOL" latex -o "$FORMAT" +"$ISABELLE_TOOL" latex -o bbl +"$ISABELLE_TOOL" latex -o "$FORMAT" +"$ISABELLE_TOOL" latex -o "$FORMAT" +"$ISABELLE_HOME/doc-src/sedindex" root +[ -f root.out ] && "$ISABELLE_HOME/doc-src/fixbookmarks" root.out +"$ISABELLE_TOOL" latex -o "$FORMAT" +"$ISABELLE_TOOL" latex -o "$FORMAT" diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/document/root.bib --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/document/root.bib Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,117 @@ +@unpublished{IsarRef, + author = "Markus Wenzel", + title = "The {Isabelle/Isar} Reference Manual", + note = "Part of the Isabelle distribution, \url{http://isabelle.in.tum.de/doc/isar-ref.pdf}." +} + +@book {Jacobson1985, + author = "Nathan Jacobson", + title = "Basic Algebra", + volume = "I", + publisher = "Freeman", + edition = "2nd", + year = 1985, + available = { CB } +} + +% TYPES 2006 + +@inproceedings{HaftmannWenzel2007, + author = "Florian Haftmann and Makarius Wenzel", + title = "Constructive Type Classes in {Isabelle}", + pages = "160--174", + crossref = "AltenkirchMcBride2007", + available = { CB } +} + +@proceedings{AltenkirchMcBride2007, + editor = "Thorsten Altenkirch and Connor McBride", + title = "Types for Proofs and Programs, TYPES 2006, Nottingham, UK", + booktitle = "Types for Proofs and Programs, TYPES 2006, Nottingham, UK", + publisher = "Springer", + series = "LNCS 4502", + year = 2007 +} + + +@techreport{Ballarin2006a, + author = "Clemens Ballarin", + title = "Interpretation of Locales in {Isabelle}: Managing Dependencies between Locales", + institution = "Technische Universit{\"a}t M{\"u}nchen", + number = "TUM-I0607", + year = 2006 +} + +% TYPES 2003 + +@inproceedings{Ballarin2004a, + author = "Clemens Ballarin", + title = "Locales and Locale Expressions in {Isabelle/Isar}", + pages = "34--50", + crossref = "BerardiEtAl2004" +} + +@proceedings{BerardiEtAl2004, + editor = "Stefano Berardi and Mario Coppo and Ferruccio Damiani", + title = "Types for Proofs and Programs, TYPES 2003, Torino, Italy", + booktitle = "Types for Proofs and Programs, TYPES 2003, Torino, Italy", + publisher = "Springer", + series = "LNCS 3085", + year = 2004 +} + +% TYPES 2008 + +@inproceedings{HaftmannWenzel2009, + author = "Florian Haftmann and Makarius Wenzel", + title = "Local theory specifications in {Isabelle}/{Isar}", + pages = "153--168", + crossref = "BerardiEtAl2009" +} + +@proceedings{BerardiEtAl2009, + editor = "Stefano Berardi and Ferruccio Damiani and Ugo de Liguoro", + title = "Types for Proofs and Programs, TYPES 2008, Torino, Italy", + booktitle = "Types for Proofs and Programs, TYPES 2008, Torino, Italy", + series = "LNCS 5497", + publisher = "Springer", + year = 2009 +} + +% MKM 2006 + +@inproceedings{Ballarin2006b, + author = "Clemens Ballarin", + title = "Interpretation of Locales in {Isabelle}: Theories and Proof Contexts", + pages = "31--43", + crossref = "BorweinFarmer2006" +} + +@proceedings{BorweinFarmer2006, + editor = "Jonathan M. Borwein and William M. Farmer", + title = "Mathematical knowledge management, MKM 2006, Wokingham, UK", + booktitle = "Mathematical knowledge management, MKM 2006, Wokingham, UK", + series = "LNCS 4108", + publisher = "Springer", + year = 2006, + available = { CB } +} + +% TPHOLs 1999 + +@inproceedings{KammullerEtAl1999, + author = "Florian Kamm{\"u}ller and Markus Wenzel and Lawrence C. Paulson", + title = "Locales: A Sectioning Concept for {Isabelle}", + pages = "149--165", + crossref = "BertotEtAl1999", + available = { CB } +} + +@book{BertotEtAl1999, + editor = "Y. Bertot and G. Dowek and A. Hirschowitz and C. Paulin and L. Th{\'e}ry", + title = "Theorem Proving in Higher Order Logics: TPHOLs'99, Nice, France", + booktitle = "Theorem Proving in Higher Order Logics: TPHOLs'99, Nice, France", + publisher = "Springer", + series = "LNCS 1690", + year = 1999 +} diff -r 75d8778f94d3 -r 54da920baf38 doc-src/Locales/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/document/root.tex Mon Aug 27 21:19:16 2012 +0200 @@ -0,0 +1,62 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{amsmath} +\usepackage{isabelle,isabellesym} +\usepackage{verbatim} +\usepackage{alltt} +\usepackage{array} + +\usepackage{amssymb} + +\usepackage{pdfsetup} + +\usepackage{ifpdf} +\ifpdf\relax\else\def\pgfsysdriver{pgfsys-dvi.def}\fi +\usepackage{tikz} +\usepackage{subfigure} + +\isadroptag{theory} +\isafoldtag{proof} + +% urls in roman style, theory text in typewriter +\urlstyle{rm} +\isabellestyle{tt} + + +\begin{document} + +\title{Tutorial to Locales and Locale Interpretation% +\thanks{Published in L.~Lamb\'an, A.~Romero, J.~Rubio, editors, {\em Contribuciones Cient\'{\i}ficas en honor de Mirian Andr\'es.} Servicio de Publicaciones de la Universidad de La Rioja, Logro\~no, Spain, 2010. Reproduced by permission.}} +\author{Clemens Ballarin} +\date{} + +\maketitle + +\begin{abstract} + Locales are Isabelle's approach for dealing with parametric + theories. They have been designed as a module system for a + theorem prover that can adequately represent the complex + inter-dependencies between structures found in abstract algebra, but + have proven fruitful also in other applications --- for example, + software verification. + + Both design and implementation of locales have evolved considerably + since Kamm\"uller did his initial experiments. Today, locales + are a simple yet powerful extension of the Isar proof language. + The present tutorial covers all major facilities of locales. It is + intended for locale novices; familiarity with Isabelle and Isar is + presumed. +\end{abstract} + +\parindent 0pt\parskip 0.5ex + +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 75d8778f94d3 -r 54da920baf38 doc-src/ROOT --- a/doc-src/ROOT Mon Aug 27 21:04:37 2012 +0200 +++ b/doc-src/ROOT Mon Aug 27 21:19:16 2012 +0200 @@ -101,13 +101,15 @@ "~~/src/HOL/Library/OptionalSugar" theories Sugar -session Locales (doc) in "Locales/Locales" = HOL + - options [browser_info = false, document = false, - document_dump = document, document_dump_mode = "tex"] +session Locales (doc) in "Locales" = HOL + + options [document_variants = "locales"] theories Examples1 Examples2 Examples3 + files + "document/build" + "document/root.tex" session Logics (doc) in "Logics" = Pure + options [document_variants = "logics"]