# HG changeset patch # User ballarin # Date 1212489262 -7200 # Node ID d1d35284542fbcb029a0b0ec3eb4b9303868a11e # Parent 9681f347b6f54cb71fa49c815dba70ba25e6d8a5 New version covering interpretation. diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/IsaMakefile --- a/doc-src/Locales/IsaMakefile Tue Jun 03 11:55:35 2008 +0200 +++ b/doc-src/Locales/IsaMakefile Tue Jun 03 12:34:22 2008 +0200 @@ -23,7 +23,8 @@ HOL: @cd $(SRC)/HOL; $(ISATOOL) make HOL -$(LOG)/HOL-Locales.gz: $(OUT)/HOL Locales/ROOT.ML Locales/Locales.thy \ +$(LOG)/HOL-Locales.gz: $(OUT)/HOL Locales/ROOT.ML Locales/Examples.thy \ + Locales/Examples1.thy Locales/Examples2.thy Locales/Examples3.thy \ Locales/document/root.tex Locales/document/root.bib @$(USEDIR) $(OUT)/HOL Locales diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/Examples.thy Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,683 @@ +(*<*) +theory Examples +imports GCD +begin + +hide const Lattices.lattice +pretty_setmargin 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 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 +\[ +%\label{eq-fact-in-context} + @{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 attributes, which + may control proof procedures. Contexts also contain syntax information + for parameters and for terms depending on them. + *} + +section {* Simple Locales *} + +text {* + Locales can be seen as persistent contexts. 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 {* The parameter of this locale is @{term le}, with infix syntax + @{text \}. There is an implicit type parameter @{typ "'a"}. It + is not necessary to declare parameter types: most general types will + be inferred from the context elements for all parameters. + + The above declaration not only introduces the locale, it also + defines the \emph{locale predicate} @{term partial_order} with + definition @{thm [source] partial_order_def}: + @{thm [display, indent=2] partial_order_def} + + 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. There are 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{fun}, \isakeyword{function} & recursive function \\ + \isakeyword{abbreviation} & syntactic abbreviation \\ + \isakeyword{theorem}, etc.\ & theorem statement with proof \\ + \isakeyword{theorems}, etc.\ & redeclaration of theorems +\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 {* A definition in a locale depends on the locale parameters. + Here, a global constant @{term partial_order.less} is declared, which is lifted over the + locale parameter @{term le}. Its definition is the global theorem + @{thm [source] partial_order.less_def}: + @{thm [display, indent=2] partial_order.less_def} + At the same time, the locale is extended by syntax information + hiding this construction in the context of the locale. That is, + @{term "partial_order.less le"} is printed and parsed as infix + @{text \}. Finally, the conclusion of the definition is added to + the locale, @{thm [locale=partial_order, source] less_def}: + @{thm [locale=partial_order, display, indent=2] less_def} + *} + +text {* As an example of a theorem statement in the locale, here is the + derivation of a transitivity law. *} + + 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, assumptions and theorems of the + locale may be used. 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. *} + +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. + + In the block below, notions of infimum and supremum together with + theorems are introduced for partial orders. + *} + + 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 {* In fact, many more theorems need to be shown for a usable + theory of partial orders. The + above two serve as illustrative examples. *} + +text {* + Two commands are provided to inspect locales: + \isakeyword{print\_locales} lists the names of all locales of the + theory; \isakeyword{print\_locale}~$n$ prints the parameters and + assumptions of locale $n$; \isakeyword{print\_locale!}~$n$ + additionally outputs the conclusions. + + The syntax of the locale commands discussed in this tutorial is + shown in Table~\ref{tab:commands}. See the + Isabelle/Isar Reference Manual~\cite{IsarRef} + for full documentation. *} + + +section {* Import *} + +text {* +\label{sec:import} + + 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. + + With locales, this inheritance is achieved through \emph{import} of a + locale. The import comes before the context elements. + *} + + locale lattice = partial_order + + assumes ex_inf: "\inf. partial_order.is_inf le x y inf" + and ex_sup: "\sup. partial_order.is_sup le x y sup" + begin + +text {* Note that the assumptions above refer to the predicates for infimum + and supremum defined in @{text partial_order}. In the current + implementation of locales, syntax from definitions of the imported + locale is unavailable in the locale declaration, neither are their + names. Hence we refer to the constants of the theory. The names + and syntax is available below, in the context of the locale. *} + + 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. + 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: + "lattice.meet le x (lattice.join le y z) = + lattice.join le (lattice.meet le x y) (lattice.meet le 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 hierachy obtained through these declarations is shown in Figure~\ref{fig:lattices}(a). + +\begin{figure} +\hrule \vspace{2ex} +\begin{center} +\subfigure[Declared hierachy]{ +\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 *} + +text {* +\label{sec:changing-the-hierarchy} + + Total orders are lattices. Hence, by deriving the lattice + axioms for total orders, the hierarchy may be changed + and @{text lattice} be placed between @{text partial_order} + and @{text total_order}, as shown in Figure~\ref{fig:lattices}(b). + Changes to the locale hierarchy may be declared + with the \isakeyword{interpretation} command. *} + + interpretation %visible total_order \ lattice + +txt {* This enters the context of locale @{text total_order}, in + which the goal @{subgoals [display]} must be shown. First, the + locale predicate needs to be unfolded --- for example using its + definition or by introduction rules + provided by the locale package. The methods @{text intro_locales} + and @{text unfold_locales} automate this. 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}, hence @{text unfold_locales} + is appropriate. *} + + proof unfold_locales + +txt {* Since both @{text lattice} and @{text total_order} + inherit @{text partial_order}, the assumptions of the latter are + discharged, and the only subgoals that remain are the assumptions + introduced in @{text lattice} @{subgoals [display]} + The proof for the first subgoal is *} + + 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 {* 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, total orders are distributive lattices. *} + + interpretation 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). *} + +(*<*) +end +(*>*) diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/Examples1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/Examples1.thy Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,94 @@ +(*<*) +theory Examples1 +imports Examples +begin + +(*>*) + +section {* Use of Locales in Theories and Proofs *} + +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}. + + The changes of the locale + hierarchy from the previous sections are examples of + interpretations. The command \isakeyword{interpretation} $l_1 + \subseteq l_2$ is said to \emph{interpret} locale $l_2$ in the + context of $l_1$. It causes all theorems of $l_2$ to be made + available in $l_1$. The interpretation is \emph{dynamic}: not only + theorems already present in $l_2$ are available in $l_1$. Theorems + that will be added to $l_2$ in future will automatically be + propagated to $l_1$. + + Locales can also be interpreted in the contexts of theories and + structured proofs. These interpretations are dynamic, too. + Theorems added to locales will be propagated to theories. + In this section the interpretation in + theories is illustrated; interpretation in proofs is analogous. + As an example consider, the type of natural numbers @{typ nat}. The + order relation @{text \} is a total order over @{typ nat}, + divisibility @{text dvd} is a distributive lattice. + + We start with the + interpretation that @{text \} is a partial order. The facilities of + the interpretation command are explored in three versions. + *} + + +subsection {* First Version: Replacement of Parameters Only *} + +text {* +\label{sec:po-first} + + In the most basic form, interpretation just replaces the locale + parameters by terms. The following command interprets the locale + @{text partial_order} in the global context of the theory. The + parameter @{term le} is replaced by @{term "op \ :: nat \ nat \ bool"}. *} + + interpretation %visible nat: partial_order ["op \ :: nat \ nat \ bool"] +txt {* The locale name is succeeded by a \emph{parameter + instantiation}. In general, this is a list of terms, which refer to + the parameters in the order of declaration in the locale. The + locale name is preceded by an optional \emph{interpretation prefix}, + which is used to qualify names from the locale in the global + context. + + The command creates the goal @{subgoals [display]} which can be shown + easily:% +\footnote{Note that @{text op} binds tighter than functions + application: parentheses around @{text "op \"} are not necessary.} + *} + by unfold_locales auto + +text {* Now theorems from the locale are available in the theory, + interpreted for natural numbers, for example @{thm [source] + nat.trans}: @{thm [display, indent=2] nat.trans} + + In order to avoid unwanted hiding of theorems, interpretation + accepts a qualifier, @{text nat} in the example, which is applied to + all names processed by the + interpretation. If present the qualifier is mandatory --- that is, + the above theorem cannot be referred to simply as @{text trans}. + *} + + +subsection {* Second Version: Replacement of Definitions *} + +text {* The above interpretation also creates the theorem + @{thm [source] nat.less_le_trans}: @{thm [display, indent=2] + nat.less_le_trans} + Here, @{term "partial_order.less (op \ :: nat \ nat \ bool)"} + represents the strict order, although @{text "<"} is defined on + @{typ nat}. Interpretation enables to map concepts + introduced in locales through definitions to the corresponding + concepts of the theory.% +\footnote{This applies not only to \isakeyword{definition} but also to + \isakeyword{inductive}, \isakeyword{fun} and \isakeyword{function}.} + *} + +(*<*) +end +(*>*) diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/Examples2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/Examples2.thy Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,34 @@ +(*<*) +theory Examples2 +imports Examples +begin + +(*>*) + +text {* This is achieved by unfolding suitable equations during + interpretation. These equations are given after the keyword + \isakeyword{where} and require proofs. The revised command, + replacing @{text "\"} by @{text "<"}, is: *} + +interpretation %visible nat: partial_order ["op \ :: [nat, nat] \ bool"] + where "partial_order.less op \ (x::nat) y = (x < y)" +proof - + txt {* The goals are @{subgoals [display]} + The proof that @{text \} is a partial order is a above. *} + show "partial_order (op \ :: nat \ nat \ bool)" + by unfold_locales auto + txt {* The second goal is shown by unfolding the + definition of @{term "partial_order.less"}. *} + show "partial_order.less op \ (x::nat) 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 a locale. + Hence, the correct interpretation of @{text + "partial_order.less_def"} is obtained manually with @{text OF}. + *} + +(*<*) +end +(*>*) \ No newline at end of file diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/Examples3.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/Examples3.thy Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,547 @@ +(*<*) +theory Examples3 +imports Examples +begin + +(*>*) + +subsection {* Third Version: Local Interpretation *} + +text {* In the above example, the fact that @{text \} is a partial + order for the natural numbers was used in the proof of the + second goal. In general, proofs of the equations may involve + theorems implied by the fact the assumptions of the instantiated + locale hold for the instantiating structure. If these theorems have + been shown abstractly in the locale they can be made available + conveniently in the context through an auxiliary local interpretation (keyword + \isakeyword{interpret}). This interpretation is inside the proof of the global + interpretation. The third revision of the example illustrates this. *} + +interpretation %visible nat: partial_order ["op \ :: nat \ nat \ bool"] + where "partial_order.less (op \) (x::nat) y = (x < y)" +proof - + show "partial_order (op \ :: nat \ nat \ bool)" + by unfold_locales auto + then interpret nat: partial_order ["op \ :: [nat, nat] \ bool"] . + show "partial_order.less (op \) (x::nat) y = (x < y)" + unfolding nat.less_def by auto +qed + +text {* The inner interpretation does not require an + elaborate new proof, it is immediate from the preceeding fact and + proved with ``.''. + This interpretation enriches the local proof context by + the very theorems also obtained in the interpretation from + Section~\ref{sec:po-first}, and @{text nat.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 closing \isakeyword{qed} --- and are + then replaced by those with the desired substitutions of the strict + order. *} + + +subsection {* Further Interpretations *} + +text {* Further interpretations are necessary to reuse theorems from + the other locales. In @{text lattice} the operations @{text \} and + @{text \} are substituted by @{term "min :: nat \ nat \ nat"} and + @{term "max :: nat \ nat \ nat"}. The entire proof for the + interpretation is reproduced in order to give an example of a more + elaborate interpretation proof. *} + +interpretation %visible nat: lattice ["op \ :: nat \ nat \ bool"] + where "lattice.meet op \ (x::nat) y = min x y" + and "lattice.join op \ (x::nat) y = max x y" +proof - + show "lattice (op \ :: nat \ nat \ bool)" + txt {* We have already shown that this is a partial order, *} + apply unfold_locales + txt {* hence only the lattice axioms remain to be shown: @{subgoals + [display]} After unfolding @{text is_inf} and @{text is_sup}, *} + apply (unfold nat.is_inf_def nat.is_sup_def) + txt {* the goals become @{subgoals [display]} which can be solved + by Presburger arithmetic. *} + by arith+ + txt {* 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 nat: lattice ["op \ :: nat \ nat \ bool"] . + show "lattice.meet op \ (x::nat) y = min x y" + by (bestsimp simp: nat.meet_def nat.is_inf_def) + show "lattice.join op \ (x::nat) y = max x y" + by (bestsimp simp: nat.join_def nat.is_sup_def) +qed + +text {* That the relation @{text \} is a total order completes this + sequence of interpretations. *} + +interpretation %visible nat: total_order ["op \ :: nat \ nat \ bool"] + by unfold_locales arith + +text {* Theorems that are available in the theory at this point are shown in + Table~\ref{tab:nat-lattice}. + +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{l} + @{thm [source] nat.less_def} from locale @{text partial_order}: \\ + \quad @{thm nat.less_def} \\ + @{thm [source] nat.meet_left} from locale @{text lattice}: \\ + \quad @{thm nat.meet_left} \\ + @{thm [source] nat.join_distr} from locale @{text distrib_lattice}: \\ + \quad @{thm nat.join_distr} \\ + @{thm [source] nat.less_total} from locale @{text total_order}: \\ + \quad @{thm nat.less_total} +\end{tabular} +\end{center} +\hrule +\caption{Interpreted theorems for @{text \} on the natural numbers.} +\label{tab:nat-lattice} +\end{table} + + Note that since the locale hierarchy reflects that total orders are + distributive lattices, an explicit interpretation of distributive + lattices for the order relation on natural numbers is not neccessary. + + Why not push this idea further and just give the last interpretation + as a single interpretation instead of the sequence of three? The + reasons for this are twofold: +\begin{itemize} +\item + Often it is easier to work in an incremental fashion, because later + interpretations require theorems provided by earlier + interpretations. +\item + Assume that a definition is made in some locale $l_1$, and that $l_2$ + imports $l_1$. Let an equation for the definition be + proved in an interpretation of $l_2$. The equation will be unfolded + in interpretations of theorems added to $l_2$ or below in the import + hierarchy, but not for theorems added above $l_2$. + Hence, an equation interpreting a definition should always be given in + an interpretation of the locale where the definition is made, not in + an interpretation of a locale further down the hierarchy. +\end{itemize} + *} + + +subsection {* Lattice @{text "dvd"} on @{typ nat} *} + +text {* Divisibility on the natural numbers is a distributive lattice + but not a total order. Interpretation again proceeds + incrementally. *} + +interpretation nat_dvd: partial_order ["op dvd :: nat \ nat \ bool"] + where "partial_order.less op dvd (x::nat) y = (x dvd y \ x \ y)" +proof - + show "partial_order (op dvd :: nat \ nat \ bool)" + by unfold_locales (auto simp: dvd_def) + then interpret nat_dvd: partial_order ["op dvd :: nat \ nat \ bool"] . + show "partial_order.less op dvd (x::nat) y = (x dvd y \ x \ y)" + apply (unfold nat_dvd.less_def) + apply auto + done +qed + +text {* Note that there is no symbol for strict divisibility. Instead, + interpretation substitutes @{term "x dvd y \ x \ y"}. *} + +interpretation nat_dvd: lattice ["op dvd :: nat \ nat \ bool"] + where nat_dvd_meet_eq: + "lattice.meet op dvd (x::nat) y = gcd (x, y)" + and nat_dvd_join_eq: + "lattice.join op dvd (x::nat) y = lcm (x, y)" +proof - + show "lattice (op dvd :: nat \ nat \ bool)" + apply unfold_locales + apply (unfold nat_dvd.is_inf_def nat_dvd.is_sup_def) + apply (rule_tac x = "gcd (x, y)" in exI) + apply auto [1] + apply (rule_tac x = "lcm (x, y)" in exI) + apply (auto intro: lcm_dvd1 lcm_dvd2 lcm_least) + done + then interpret nat_dvd: lattice ["op dvd :: nat \ nat \ bool"] . + show "lattice.meet op dvd (x::nat) y = gcd (x, y)" + apply (unfold nat_dvd.meet_def) + apply (rule the_equality) + apply (unfold nat_dvd.is_inf_def) + by auto + show "lattice.join op dvd (x::nat) y = lcm (x, y)" + apply (unfold nat_dvd.join_def) + apply (rule the_equality) + apply (unfold nat_dvd.is_sup_def) + by (auto intro: lcm_dvd1 lcm_dvd2 lcm_least) +qed + +text {* Equations @{thm [source] nat_dvd_meet_eq} and @{thm [source] + nat_dvd_join_eq} are named since they are handy in the proof of + the subsequent interpretation. *} + +ML %invisible {* set quick_and_dirty *} + +(* +definition + is_lcm :: "nat \ nat \ nat \ bool" where + "is_lcm p m n \ m dvd p \ n dvd p \ + (\d. m dvd d \ n dvd d \ p dvd d)" + +lemma is_gcd: "is_lcm (lcm (m, n)) m n" + by (simp add: is_lcm_def lcm_least) + +lemma gcd_lcm_distr_lemma: + "[| is_gcd g1 x l1; is_lcm l1 y z; is_gcd g2 x y; is_gcd g3 x z |] ==> is_lcm g1 g2 g3" +apply (unfold is_gcd_def is_lcm_def dvd_def) +apply (clarsimp simp: mult_ac) +apply (blast intro: mult_is_0) +thm mult_is_0 [THEN iffD1] +*) + +lemma %invisible gcd_lcm_distr: + "gcd (x, lcm (y, z)) = lcm (gcd (x, y), gcd (x, z))" + sorry + +ML %invisible {* reset quick_and_dirty *} + +interpretation %visible nat_dvd: + distrib_lattice ["op dvd :: nat \ nat \ bool"] + apply unfold_locales + txt {* @{subgoals [display]} *} + apply (unfold nat_dvd_meet_eq nat_dvd_join_eq) + txt {* @{subgoals [display]} *} + apply (rule gcd_lcm_distr) done + + +text {* Theorems that are available in the theory after these + interpretations are shown in Table~\ref{tab:nat-dvd-lattice}. + +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{l} + @{thm [source] nat_dvd.less_def} from locale @{text partial_order}: \\ + \quad @{thm nat_dvd.less_def} \\ + @{thm [source] nat_dvd.meet_left} from locale @{text lattice}: \\ + \quad @{thm nat_dvd.meet_left} \\ + @{thm [source] nat_dvd.join_distr} from locale @{text distrib_lattice}: \\ + \quad @{thm nat_dvd.join_distr} \\ +\end{tabular} +\end{center} +\hrule +\caption{Interpreted theorems for @{text dvd} on the natural numbers.} +\label{tab:nat-dvd-lattice} +\end{table} + *} + +text {* + The full syntax of the interpretation commands is shown in + Table~\ref{tab:commands}. The grammar refers to + \textit{expr}, which stands for a \emph{locale} expression. Locale + expressions are discussed in Section~\ref{sec:expressions}. + *} + + +section {* Locale Expressions *} + +text {* +\label{sec:expressions} + + 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. + + Inspecting the grammar of locale commands in + Table~\ref{tab:commands} reveals that the import of a locale can be + more than just a single locale. In general, the import is a + \emph{locale expression}. Locale expressions enable to combine locales + and rename parameters. A locale name is a locale expression. If + $e_1$ and $e_2$ are locale expressions then $e_1 + e_2$ is their + \emph{merge}. If $e$ is an expression, then $e\:q_1 \ldots q_n$ is + a \emph{renamed expression} where the parameters in $e$ are renamed + to $q_1 \ldots q_n$. Using a locale expression, a locale for order + preserving maps can be declared in the following way. *} + + locale order_preserving = + partial_order + partial_order le' (infixl "\" 50) + + fixes \ :: "'a \ 'b" + assumes hom_le: "x \ y \ \ x \ \ y" + +text {* The second line contains the expression, which is the + merge of two partial order locales. The parameter of the second one + is @{text le'} with new infix syntax @{text \}. The + parameters of the entire locale are @{text le}, @{text le'} and + @{text \}. This is their \emph{canonical order}, + which is obtained by a left-to-right traversal of the expression, + where only the new parameters are appended to the end of the list. The + parameters introduced in the locale elements of the declaration + follow. + + In renamings parameters are referred to by position in the canonical + order; an underscore is used to skip a parameter position, which is + then not renamed. Renaming deletes the syntax of a parameter unless + a new mixfix annotation is given. + + Parameter renamings are morphisms between locales. These can be + lifted to terms and theorems and thus be applied to assumptions and + conclusions. The assumption of a merge is the conjunction of the + assumptions of the merged locale. The conclusions of a merge are + obtained by appending the conclusions of the left locale and of the + right locale. *} + +text {* The locale @{text order_preserving} contains theorems for both + orders @{text \} and @{text \}. How can one refer to a theorem for + a particular order, @{text \} or @{text \}? Names in locales are + qualified by the locale parameters. More precisely, a name is + qualified by the parameters of the locale in which its declaration + occurs. Here are examples: *} + +context %invisible order_preserving begin + +text {* + @{thm [source] le.less_le_trans}: @{thm le.less_le_trans} + + @{thm [source] le_le'_\.hom_le}: @{thm le_le'_\.hom_le} + *} + +text {* When renaming a locale, the morphism is also applied + to the qualifiers. Hence theorems for the partial order @{text \} + are qualified by @{text le'}. For example, @{thm [source] + le'.less_le_trans}: @{thm [display, indent=2] le'.less_le_trans} *} + +end %invisible + +text {* This example reveals that there is no infix syntax for the strict + version of @{text \}! This can, of course, not be introduced + automatically, but it can be declared manually through an abbreviation. + *} + + abbreviation (in order_preserving) + less' (infixl "\" 50) where "less' \ partial_order.less le'" + +text {* Now the theorem is displayed nicely as + @{thm [locale=order_preserving] le'.less_le_trans}. *} + +text {* Not only names of theorems are qualified. In fact, all names + are qualified, in particular names introduced by definitions and + abbreviations. The name of the strict order of @{text \} is @{text + le.less} and therefore @{text le'.less} is the name of the strict + order of @{text \}. Hence, the equation in the above abbreviation + could have been written as @{text "less' \ le'.less"}. *} + +text {* Two more locales illustrate working with locale expressions. + A map @{text \} is a lattice homomorphism if it preserves meet and join. *} + + locale lattice_hom = lattice + lattice le' (infixl "\" 50) + + fixes \ + assumes hom_meet: + "\ (lattice.meet le x y) = lattice.meet le' (\ x) (\ y)" + and hom_join: + "\ (lattice.join le x y) = lattice.join le' (\ x) (\ y)" + + abbreviation (in lattice_hom) + meet' (infixl "\''" 50) where "meet' \ le'.meet" + abbreviation (in lattice_hom) + join' (infixl "\''" 50) where "join' \ le'.join" + +text {* A homomorphism is an endomorphism if both orders coincide. *} + + locale lattice_end = + lattice_hom le (infixl "\" 50) le (infixl "\" 50) + +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. Renamings are + indicated by $\sqsubseteq \mapsto \preceq$ etc. The expression + imported by @{text lattice_end} identifies the first and second + parameter of @{text lattice_hom}. 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. This is not the + case! Inheritance paths with identical morphisms are detected and + the conclusions of the respecitve 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. *} + + interpretation 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 {* Theorems and other declarations --- syntax, in particular --- + from the locale @{text order_preserving} are now active in @{text + lattice_hom}, for example + + @{thm [locale=lattice_hom, source] le'.less_le_trans}: + @{thm [locale=lattice_hom] le'.less_le_trans} + *} + + + +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, we + 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 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 original work of Kamm\"uller on locales \cite{KammullerEtAl1999} + may be of interest from a historical perspective. The mathematical + background on orders and lattices is taken from Jacobson's textbook + on algebra \cite[Chapter~8]{Jacobson1985}. + *} + +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} \\[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{rename} & ::= + & \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' \\ + \textit{expr} & ::= + & \textit{renamed-expr} ( ``\textbf{+}'' \textit{renamed-expr} )$^*$ \\ + \textit{renamed-expr} & ::= + & ( \textit{qualified-name} $|$ + ``\textbf{(}'' \textit{expr} ``\textbf{)}'' ) \textit{rename}$^*$ \\[2ex] + + \multicolumn{3}{l}{Declaration of Locales} \\ + + \textit{locale} & ::= + & \textit{element}$^+$ \\ + & | & \textit{locale-expr} [ ``\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{insts} & ::= & [ ``\textbf{[}'' \textit{term}$^+$ + ``\textbf{]}'' ] \\ + & & [ \textbf{where} \textit{equation} ( \textbf{and} + \textit{equation} )$^*$ ] \\ + \textit{toplevel} & ::= + & \textbf{interpretation} \textit{name} ( ``$<$'' $|$ + ``$\subseteq$'' ) \textit{expr} \textit{proof} \\ + & | + & \textbf{interpretation} [ \textit{attr-name} ``\textbf{:}'' ] + \textit{expr} \textit{insts} \textit{proof} \\ + & | + & \textbf{interpret} [ \textit{attr-name} ``\textbf{:}'' ] + \textit{expr} \textit{insts} \textit{proof} \\[2ex] + + \multicolumn{3}{l}{Diagnostics} \\ + + \textit{toplevel} & ::= + & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{locale} \\ + & | & \textbf{print\_locales} +\end{tabular} +\end{center} +\hrule +\caption{Syntax of Locale Commands.} +\label{tab:commands} +\end{table} + *} + +text {* \textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, + Christian Sternagel and Makarius Wenzel have made useful comments on + a draft of this document. *} + +(*<*) +end +(*>*) diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/Locales.thy --- a/doc-src/Locales/Locales/Locales.thy Tue Jun 03 11:55:35 2008 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1172 +0,0 @@ -(* - Title: Locales in Isabelle/Isar - Id: $Id$ - Author: Clemens Ballarin, started 31 January 2003 - Copyright (c) 2003 by Clemens Ballarin -*) - -(*<*) -theory Locales imports Main begin - -ML {* print_mode := "type_brackets" :: !print_mode; *} -(*>*) - -section {* Overview *} - -text {* - The present text is based on~\cite{Ballarin2004a}. It was updated - for for Isabelle2005, but does not cover locale interpretation. - - Locales are an extension of the Isabelle proof assistant. They - provide support for modular reasoning. Locales were initially - developed by Kamm\"uller~\cite{Kammuller2000} to support reasoning - in abstract algebra, but are applied also in other domains --- for - example, bytecode verification~\cite{Klein2003}. - - Kamm\"uller's original design, implemented in Isabelle99, provides, in - addition to - means for declaring locales, a set of ML functions that were used - along with ML tactics in a proof. In the meantime, the input format - for proof in Isabelle has changed and users write proof - scripts in ML only rarely if at all. Two new proof styles are - available, and can - be used interchangeably: linear proof scripts that closely resemble ML - tactics, and the structured Isar proof language by - Wenzel~\cite{Wenzel2002a}. Subsequently, Wenzel re-implemented - locales for - the new proof format. The implementation, available with - Isabelle2002 or later, constitutes a complete re-design and exploits that - both Isar and locales are based on the notion of context, - and thus locales are seen as a natural extension of Isar. - Nevertheless, locales can also be used with proof scripts: - their use does not require a deep understanding of the structured - Isar proof style. - - At the same time, Wenzel considerably extended locales. The most - important addition are locale expressions, which allow to combine - locales more freely. Previously only - linear inheritance was possible. Now locales support multiple - inheritance through a normalisation algorithm. New are also - structures, which provide special syntax for locale parameters that - represent algebraic structures. - - Unfortunately, Wenzel provided only an implementation but hardly any - documentation. Besides providing documentation, the present paper - is a high-level description of locales, and in particular locale - expressions. It is meant as a first step towards the semantics of - locales, and also as a base for comparing locales with module concepts - in other provers. It also constitutes the base for future - extensions of locales in Isabelle. - The description was derived mainly by experimenting - with locales and partially also by inspecting the code. - - The main contribution of the author of the present paper is the - abstract description of Wenzel's version of locales, and in - particular of the normalisation algorithm for locale expressions (see - Section~\ref{sec-normal-forms}). Contributions to the - implementation are confined to bug fixes and to provisions that - enable the use of locales with linear proof scripts. - - Concepts are introduced along with examples, so that the text can be - used as tutorial. It is assumed that the reader is somewhat - familiar with Isabelle proof scripts. Examples have been phrased as - structured - Isar proofs. However, in order to understand the key concepts, - including locales expressions and their normalisation, detailed - knowledge of Isabelle is not necessary. - -\nocite{Nipkow2003,Wenzel2002b,Wenzel2003} -*} - -section {* Locales: Beyond Proof Contexts *} - -text {* - In tactic-based provers the application of a sequence of proof - tactics leads to a proof state. This state is usually hard to - predict from looking at the tactic script, unless one replays the - proof step-by-step. The structured proof language Isar is - different. It is additionally based on \emph{proof contexts}, - which are directly visible in Isar scripts, and since tactic - sequences tend to be short, this commonly leads to clearer proof - scripts. - - Goals are stated with the \textbf{theorem} - command. This is followed by a proof. When discharging a goal - requires an elaborate argument - (rather than the application of a single tactic) a new context - may be entered (\textbf{proof}). Inside the context, variables may - be fixed (\textbf{fix}), assumptions made (\textbf{assume}) and - intermediate goals stated (\textbf{have}) and proved. The - assumptions must be dischargeable by premises of the surrounding - goal, and once this goal has been proved (\textbf{show}) the proof context - can be closed (\textbf{qed}). Contexts inherit from surrounding - contexts, but it is not possible - to export from them (with exception of the proved goal); - they ``disappear'' after the closing \textbf{qed}. - Facts may have attributes --- for example, identifying them as - default to the simplifier or classical reasoner. - - Locales extend proof contexts in various ways: - \begin{itemize} - \item - Locales are usually \emph{named}. This makes them persistent. - \item - Fixed variables may have \emph{syntax}. - \item - It is possible to \emph{add} and \emph{export} facts. - \item - Locales can be combined and modified with \emph{locale - expressions}. - \end{itemize} - The Locales facility extends the Isar language: it provides new ways - of stating and managing facts, but it does not modify the language - for proofs. Its purpose is to support writing modular proofs. -*} - -section {* Simple Locales *} - -subsection {* Syntax and Terminology *} - -text {* - The grammar of Isar is extended by commands for locales as shown in - Figure~\ref{fig-grammar}. - A key concept, introduced by Wenzel, is that - locales are (internally) lists - of \emph{context elements}. There are five kinds, identified - by the keywords \textbf{fixes}, \textbf{constrains}, - \textbf{assumes}, \textbf{defines} and \textbf{notes}. - - \begin{figure} - \hrule - \vspace{2ex} - \begin{small} - \begin{tabular}{l>$c<$l} - \textit{attr-name} & ::= - & \textit{name} $|$ \textit{attribute} $|$ - \textit{name} \textit{attribute} \\ - - \textit{locale-expr} & ::= - & \textit{locale-expr1} ( ``\textbf{+}'' \textit{locale-expr1} )$^*$ \\ - \textit{locale-expr1} & ::= - & ( \textit{qualified-name} $|$ - ``\textbf{(}'' \textit{locale-expr} ``\textbf{)}'' ) \\ - & & ( \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' )$^*$ \\ - - \textit{fixes} & ::= - & \textit{name} [ ``\textbf{::}'' \textit{type} ] - [ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ - \textit{mixfix} ] \\ - \textit{constrains} & ::= - & \textit{name} ``\textbf{::}'' \textit{type} \\ - \textit{assumes} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ - \textit{defines} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ - \textit{notes} & ::= - & [ \textit{attr-name} ``\textbf{=}'' ] - ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ - - \textit{element} & ::= - & \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ - & | - & \textbf{constrains} \textit{constrains} - ( \textbf{and} \textit{constrains} )$^*$ \\ - & | - & \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\ - & | - & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ - & | - & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ - \textit{element1} & ::= - & \textit{element} \\ - & | & \textbf{includes} \textit{locale-expr} \\ - - \textit{locale} & ::= - & \textit{element}$^+$ \\ - & | & \textit{locale-expr} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ - - \textit{in-target} & ::= - & ``\textbf{(}'' \textbf{in} \textit{qualified-name} ``\textbf{)}'' \\ - - \textit{theorem} & ::= & ( \textbf{theorem} $|$ \textbf{lemma} $|$ - \textbf{corollary} ) [ \textit{in-target} ] [ \textit{attr-name} ] \\ - - \textit{theory-level} & ::= & \ldots \\ - & | & \textbf{locale} \textit{name} [ ``\textbf{=}'' - \textit{locale} ] \\ - % note: legacy "locale (open)" omitted. - & | & ( \textbf{theorems} $|$ \textbf{lemmas} ) \\ - & & [ \textit{in-target} ] [ \textit{attr-name} ``\textbf{=}'' ] - ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ - & | & \textbf{declare} [ \textit{in-target} ] ( \textit{qualified-name} - [ \textit{attribute} ] )$^+$ \\ - & | & \textit{theorem} \textit{proposition} \textit{proof} \\ - & | & \textit{theorem} \textit{element1}$^*$ - \textbf{shows} \textit{proposition} \textit{proof} \\ - & | & \textbf{print\_locale} \textit{locale} \\ - & | & \textbf{print\_locales} - \end{tabular} - \end{small} - \vspace{2ex} - \hrule - \caption{Locales extend the grammar of Isar.} - \label{fig-grammar} - \end{figure} - - At the theory level --- that is, at the outer syntactic level of an - Isabelle input file --- \textbf{locale} declares a named - locale. Other kinds of locales, - locale expressions and unnamed locales, will be introduced later. When - declaring a named locale, it is possible to \emph{import} another - named locale, or indeed several ones by importing a locale - expression. The second part of the declaration, also optional, - consists of a number of context element declarations. - - A number of Isar commands have an additional, optional \emph{target} - argument, which always refers to a named locale. These commands - are \textbf{theorem} (together with \textbf{lemma} and - \textbf{corollary}), \textbf{theorems} (and - \textbf{lemmas}), and \textbf{declare}. The effect of specifying a target is - that these commands focus on the specified locale, not the - surrounding theory. Commands that are used to - prove new theorems will add them not to the theory, but to the - locale. Similarly, \textbf{declare} modifies attributes of theorems - that belong to the specified target. Additionally, for - \textbf{theorem} (and related commands), theorems stored in the target - can be used in the associated proof scripts. - - The Locales package provides a \emph{long goals format} for - propositions stated with \textbf{theorem} (and friends). While - normally a goal is just a formula, a long goal is a list of context - elements, followed by the keyword \textbf{shows}, followed by the - formula. Roughly speaking, the context elements are - (additional) premises. For an example, see - Section~\ref{sec-includes}. The list of context elements in a long goal - is also called \emph{unnamed locale}. - - Finally, there are two commands to inspect locales when working in - interactive mode: \textbf{print\_locales} prints the names of all - targets - visible in the current theory, \textbf{print\_locale} outputs the - elements of a named locale or locale expression. - - 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. -*} - -subsection {* Parameters, Assumptions and Facts *} - -text {* - From a logical point of view a \emph{context} is a formula schema of - the form -\[ - @{text "\x\<^sub>1\x\<^sub>n. \ C\<^sub>1; \ ;C\<^sub>m \ \ \"} -\] - The variables $@{text "x\<^sub>1"}, \ldots, @{text "x\<^sub>n"}$ are - called \emph{parameters}, the premises $@{text "C\<^sub>1"}, \ldots, - @{text "C\<^sub>n"}$ \emph{assumptions}. A formula @{text "F"} - holds in this context if -\begin{equation} -\label{eq-fact-in-context} - @{text "\x\<^sub>1\x\<^sub>n. \ C\<^sub>1; \ ;C\<^sub>m \ \ F"} -\end{equation} - is valid. The formula is called a \emph{fact} of the context. - - A locale allows fixing the parameters @{text - "x\<^sub>1, \, x\<^sub>n"} and making the assumptions @{text - "C\<^sub>1, \, C\<^sub>m"}. This implicitly builds the context in - which the formula @{text "F"} can be established. - Parameters of a locale correspond to the context element - \textbf{fixes}, and assumptions may be declared with - \textbf{assumes}. Using these context elements one can define - the specification of semigroups. -*} - -locale semi = - fixes prod :: "['a, 'a] \ 'a" (infixl "\" 70) - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - -text {* - The parameter @{term prod} has a - syntax annotation enabling the infix ``@{text "\"}'' in the - assumption of associativity. Parameters may have arbitrary mixfix - syntax, like constants. In the example, the type of @{term prod} is - specified explicitly. This is not necessary. If no type is - specified, a most general type is inferred simultaneously for all - parameters, taking into account all assumptions (and type - specifications of parameters, if present).% -\footnote{Type inference also takes into account type constraints, - definitions and import, as introduced later.} - - Free variables in assumptions are implicitly universally quantified, - unless they are parameters. Hence the context defined by the locale - @{term "semi"} is -\[ - @{text "\prod. \ \x y z. prod (prod x y) z = prod x (prod y z) - \ \ \"} -\] - The locale can be extended to commutative semigroups. -*} - -locale comm_semi = semi + - assumes comm: "x \ y = y \ x" - -text {* - This locale \emph{imports} all elements of @{term "semi"}. The - latter locale is called the import of @{term "comm_semi"}. The - definition adds commutativity, hence its context is -\begin{align*% -} - @{text "\prod. \ "} & - @{text "\x y z. prod (prod x y) z = prod x (prod y z);"} \\ - & @{text "\x y. prod x y = prod y x \ \ \"} -\end{align*% -} - One may now derive facts --- for example, left-commutativity --- in - the context of @{term "comm_semi"} by specifying this locale as - target, and by referring to the names of the assumptions @{text - assoc} and @{text comm} in the proof. -*} - -theorem (in comm_semi) lcomm: - "x \ (y \ z) = y \ (x \ z)" -proof - - have "x \ (y \ z) = (x \ y) \ z" by (simp add: assoc) - also have "\ = (y \ x) \ z" by (simp add: comm) - also have "\ = y \ (x \ z)" by (simp add: assoc) - finally show ?thesis . -qed - -text {* - In this equational Isar proof, ``@{text "\"}'' refers to the - right hand side of the preceding equation. - After the proof is finished, the fact @{text "lcomm"} is added to - the locale @{term "comm_semi"}. This is done by adding a - \textbf{notes} element to the internal representation of the locale, - as explained the next section. -*} - -subsection {* Locale Predicates and the Internal Representation of - Locales *} - -text {* -\label{sec-locale-predicates} - In mathematical texts, often arbitrary but fixed objects with - certain properties are considered --- for instance, an arbitrary but - fixed group $G$ --- with the purpose of establishing facts valid for - any group. These facts are subsequently used on other objects that - also have these properties. - - Locales permit the same style of reasoning. Exporting a fact $F$ - generalises the fixed parameters and leads to a (valid) formula of the - form of equation~(\ref{eq-fact-in-context}). If a locale has many - assumptions - (possibly accumulated through a number of imports) this formula can - become large and cumbersome. Therefore, Wenzel introduced - predicates that abbreviate the assumptions of locales. These - predicates are not confined to the locale but are visible in the - surrounding theory. - - The definition of the locale @{term semi} generates the \emph{locale - predicate} @{term semi} over the type of the parameter @{term prod}, - hence the predicate's type is @{typ "(['a, 'a] \ 'a) - \ bool"}. Its - definition is -\begin{equation} - \tag*{@{thm [source] semi_def}:} @{thm semi_def}. -\end{equation} - In the case where the locale has no import, the generated - predicate abbreviates all assumptions and is over the parameters - that occur in these assumptions. - - The situation is more complicated when a locale extends - another locale, as is the case for @{term comm_semi}. Two - predicates are defined. The predicate - @{term comm_semi_axioms} corresponds to the new assumptions and is - called \emph{delta predicate}, the locale - predicate @{term comm_semi} captures the content of all the locale, - including the import. - If a locale has neither assumptions nor import, no predicate is - defined. If a locale has import but no assumptions, only the locale - predicate is defined. -*} -(*<*) -ML {* - val [comm_semi_ax1, comm_semi_ax2] = thms "comm_semi.axioms"; - bind_thm ("comm_semi_ax1", comm_semi_ax1); - bind_thm ("comm_semi_ax2", comm_semi_ax2); -*} -(*>*) -text {* - The Locales package generates a number of theorems for locale and - delta predicates. All predicates have a definition and an - introduction rule. Locale predicates that are defined in terms of - other predicates (which is the case if and only if the locale has - import) also have a number of elimination rules (called - \emph{axioms}). All generated theorems for the predicates of the - locales @{text semi} and @{text comm_semi} are shown in - Figures~\ref{fig-theorems-semi} and~\ref{fig-theorems-comm-semi}, - respectively. - \begin{figure} - \hrule - \vspace{2ex} - Theorems generated for the predicate @{term semi}. - \begin{gather} - \tag*{@{thm [source] semi_def}:} @{thm semi_def} \\ - \tag*{@{thm [source] semi.intro}:} @{thm semi.intro} - \end{gather} - \hrule - \caption{Theorems for the locale predicate @{term "semi"}.} - \label{fig-theorems-semi} - \end{figure} - - \begin{figure}[h] - \hrule - \vspace{2ex} - Theorems generated for the predicate @{term "comm_semi_axioms"}. - \begin{gather} - \tag*{@{thm [source] comm_semi_axioms_def}:} @{thm - comm_semi_axioms_def} \\ - \tag*{@{thm [source] comm_semi_axioms.intro}:} @{thm - comm_semi_axioms.intro} - \end{gather} - Theorems generated for the predicate @{term "comm_semi"}. - \begin{gather} - \tag*{@{thm [source] comm_semi_def}:} @{thm - comm_semi_def} \\ - \tag*{@{thm [source] comm_semi.intro}:} @{thm - comm_semi.intro} \\ - \tag*{@{thm [source] comm_semi.axioms}:} \mbox{} \\ - \notag @{thm comm_semi_ax1} \\ - \notag @{thm comm_semi_ax2} - \end{gather} - \hrule - \caption{Theorems for the predicates @{term "comm_semi_axioms"} and - @{term "comm_semi"}.} - \label{fig-theorems-comm-semi} - \end{figure} - Note that the theorems generated by a locale - definition may be inspected immediately after the definition in the - Proof General interface \cite{Aspinall2000} of Isabelle through - the menu item ``Isabelle/Isar$>$Show me $\ldots>$Theorems''. - - Locale and delta predicates are used also in the internal - representation of locales as lists of context elements. While all - \textbf{fixes} in a - declaration generate internal \textbf{fixes}, all assumptions - of one locale declaration contribute to one internal \textbf{assumes} - element. The internal representation of @{term semi} is -\[ -\begin{array}{ll} - \textbf{fixes} & @{text prod} :: @{typ [quotes] "['a, 'a] \ 'a"} - (\textbf{infixl} @{text [quotes] "\"} 70) \\ - \textbf{assumes} & @{text [quotes] "semi prod"} \\ - \textbf{notes} & @{text assoc}: @{text [quotes] "?x \ ?y \ ?z = ?x \ (?y \ - ?z)"} -\end{array} -\] - and the internal representation of @{text [quotes] comm_semi} is -\begin{equation} -\label{eq-comm-semi} -\begin{array}{ll} - \textbf{fixes} & @{text prod} :: @{typ [quotes] "['a, 'a] \ 'a"} - ~(\textbf{infixl}~@{text [quotes] "\"}~70) \\ - \textbf{assumes} & @{text [quotes] "semi prod"} \\ - \textbf{notes} & @{text assoc}: @{text [quotes] "?x \ ?y \ ?z = ?x \ (?y \ - ?z)"} \\ - \textbf{assumes} & @{text [quotes] "comm_semi_axioms prod"} \\ - \textbf{notes} & @{text comm}: @{text [quotes] "?x \ ?y = ?y \ ?x"} \\ - \textbf{notes} & @{text lcomm}: @{text [quotes] "?x \ (?y \ ?z) = ?y \ (?x \ - ?z)"} -\end{array} -\end{equation} - The \textbf{notes} elements store facts of the - locales. The facts @{text assoc} and @{text comm} were added - during the declaration of the locales. They stem from assumptions, - which are trivially facts. The fact @{text lcomm} was - added later, after finishing the proof in the respective - \textbf{theorem} command above. - - By using \textbf{notes} in a declaration, facts can be added - to a locale directly. Of course, these must be theorems. - Typical use of this feature includes adding theorems that are not - usually used as a default rewrite rules by the simplifier to the - simpset of the locale by a \textbf{notes} element with the attribute - @{text "[simp]"}. This way it is also possible to add specialised - versions of - theorems to a locale by instantiating locale parameters for unknowns - or locale assumptions for premises. -*} - -subsection {* Definitions *} - -text {* - Definitions were available in Kamm\"uller's version of Locales, and - they are in Wenzel's. - The context element \textbf{defines} adds a definition of the form - @{text "p x\<^sub>1 \ x\<^sub>n \ t"} as an assumption, where @{term p} is a - parameter of the locale (possibly an imported parameter), and @{text - t} a term that may contain the @{text "x\<^sub>i"}. The parameter may - neither occur in a previous \textbf{assumes} or \textbf{defines} - element, nor on the right hand side of the definition. Hence - recursion is not allowed. - The parameter may, however, occur in subsequent \textbf{assumes} and - on the right hand side of subsequent \textbf{defines}. We call - @{term p} \emph{defined parameter}. -*} - -locale semi2 = semi + - fixes rprod (infixl "\" 70) - defines rprod_def: "rprod x y \ y \ x " - -text {* - This locale extends @{term semi} by a second binary operation @{text - [quotes] \} that is like @{text [quotes] \} but with - reversed arguments. The - definition of the locale generates the predicate @{term semi2}, - which is equivalent to @{text semi}, but no @{term "semi2_axioms"}. - The difference between \textbf{assumes} and \textbf{defines} lies in - the way parameters are treated on export. -*} - -subsection {* Export *} - -text {* - A fact is exported out - of a locale by generalising over the parameters and adding - assumptions as premises. For brevity of the exported theorems, - locale predicates are used. Exported facts are referenced by - writing qualified names consisting of the locale name and the fact name --- - for example, -\begin{equation} - \tag*{@{thm [source] semi.assoc}:} @{thm semi.assoc}. -\end{equation} - Defined parameters receive special treatment. Instead of adding a - premise for the definition, the definition is unfolded in the - exported theorem. In order to illustrate this we prove that the - reverse operation @{text [quotes] \} defined in the locale - @{text semi2} is also associative. -*} - -theorem (in semi2) r_assoc: "(x \ y) \ z = x \ (y \ z)" - by (simp only: rprod_def assoc) - -text {* - The exported fact is -\begin{equation} - \tag*{@{thm [source] semi2.r_assoc}:} @{thm semi2.r_assoc}. -\end{equation} - The defined parameter is not present but is replaced by its - definition. Note that the definition itself is not exported, hence - there is no @{text "semi2.rprod_def"}.% -\footnote{The definition could alternatively be exported using a - let-construct if there was one in Isabelle's meta-logic. Let is - usually defined in object-logics.} -*} - -section {* Locale Expressions *} - -text {* - Locale expressions provide a simple language for combining - locales. They are an effective means of building complex - specifications from simple ones. Locale expressions are the main - innovation of the version of Locales discussed here. Locale - expressions are also reason for introducing locale predicates. -*} - -subsection {* Rename and Merge *} - -text {* - The grammar of locale expressions is part of the grammar in - Figure~\ref{fig-grammar}. Locale names are locale - expressions, and - further expressions are obtained by \emph{rename} and \emph{merge}. -\begin{description} -\item[Rename.] - The locale expression $e\: q_1 \ldots q_n$ denotes - the locale of $e$ where pa\-ra\-me\-ters, in the order in - which they are fixed, are renamed to $q_1$ to $q_n$. - The expression is only well-formed if $n$ does not - exceed the number of parameters of $e$. Underscores denote - parameters that are not renamed. - Renaming by default removes mixfix syntax, but new syntax may be - specified. -\item[Merge.] - The locale expression $e_1 + e_2$ denotes - the locale obtained by merging the locales of $e_1$ - and $e_2$. This locale contains the context elements - of $e_1$, followed by the context elements of $e_2$. - - In actual fact, the semantics of the merge operation - is more complicated. If $e_1$ and $e_2$ are expressions containing - the same name, followed by - identical parameter lists, then the merge of both will contain - the elements of those locales only once. Details are explained in - Section~\ref{sec-normal-forms} below. - - The merge operation is associative but not commutative. The latter - is because parameters of $e_1$ appear - before parameters of $e_2$ in the composite expression. -\end{description} - - Rename can be used if a different parameter name seems more - appropriate --- for example, when moving from groups to rings, a - parameter @{text G} representing the group could be changed to - @{text R}. Besides of this stylistic use, renaming is important in - combination with merge. Both operations are used in the following - specification of semigroup homomorphisms. -*} - -locale semi_hom = comm_semi sum (infixl "\" 65) + comm_semi + - fixes hom - assumes hom: "hom (x \ y) = hom x \ hom y" - -text {* - This locale defines a context with three parameters @{text "sum"}, - @{text "prod"} and @{text "hom"}. The first two parameters have - mixfix syntax. They are associative operations, - the first of type @{typeof [locale=semi_hom] prod}, the second of - type @{typeof [locale=semi_hom] sum}. - - How are facts that are imported via a locale expression identified? - Facts are always introduced in a named locale (either in the - locale's declaration, or by using the locale as target in - \textbf{theorem}), and their names are - qualified by the parameter names of this locale. - Hence the full name of associativity in @{text "semi"} is - @{text "prod.assoc"}. Renaming parameters of a target also renames - the qualifier of facts. Hence, associativity of @{text "sum"} is - @{text "sum.assoc"}. Several parameters are separated by - underscores in qualifiers. For example, the full name of the fact - @{text "hom"} in the locale @{text "semi_hom"} is @{text - "sum_prod_hom.hom"}. - - The following example is quite artificial, it illustrates the use of - facts, though. -*} - -theorem (in semi_hom) "hom x \ (hom y \ hom z) = hom (x \ (y \ z))" -proof - - have "hom x \ (hom y \ hom z) = hom y \ (hom x \ hom z)" - by (simp add: prod.lcomm) - also have "\ = hom (y \ (x \ z))" by (simp add: hom) - also have "\ = hom (x \ (y \ z))" by (simp add: sum.lcomm) - finally show ?thesis . -qed - -text {* - Importing via a locale expression imports all facts of - the imported locales, hence both @{text "sum.lcomm"} and @{text - "prod.lcomm"} are - available in @{text "hom_semi"}. The import is dynamic --- that is, - whenever facts are added to a locale, they automatically - become available in subsequent \textbf{theorem} commands that use - the locale as a target, or a locale importing the locale. -*} - -subsection {* Normal Forms *} - -text_raw {* -\label{sec-normal-forms} -\newcommand{\I}{\mathcal{I}} -\newcommand{\F}{\mathcal{F}} -\newcommand{\N}{\mathcal{N}} -\newcommand{\C}{\mathcal{C}} -\newcommand{\App}{\mathbin{\overline{@}}} -*} - -text {* - Locale expressions are interpreted in a two-step process. First, an - expression is normalised, then it is converted to a list of context - elements. - - Normal forms are based on \textbf{locale} declarations. These - consist of an import section followed by a list of context - elements. Let $\I(l)$ denote the locale expression imported by - locale $l$. If $l$ has no import then $\I(l) = \varepsilon$. - Likewise, let $\F(l)$ denote the list of context elements, also - called the \emph{context fragment} of $l$. Note that $\F(l)$ - contains only those context elements that are stated in the - declaration of $l$, not imported ones. - -\paragraph{Example 1.} Consider the locales @{text semi} and @{text - "comm_semi"}. We have $\I(@{text semi}) = \varepsilon$ and - $\I(@{text "comm_semi"}) = @{text semi}$, and the context fragments - are -\begin{align*% -} - \F(@{text semi}) & = \left[ -\begin{array}{ll} - \textbf{fixes} & @{text prod} :: @{typ [quotes] "['a, 'a] \ 'a"} - ~(\textbf{infixl}~@{text [quotes] "\"}~70) \\ - \textbf{assumes} & @{text [quotes] "semi prod"} \\ - \textbf{notes} & @{text assoc}: @{text [quotes]"?x \ ?y \ ?z = ?x \ (?y \ - ?z)"} -\end{array} \right], \\ - \F(@{text "comm_semi"}) & = \left[ -\begin{array}{ll} - \textbf{assumes} & @{text [quotes] "comm_semi_axioms prod"} \\ - \textbf{notes} & @{text comm}: @{text [quotes] "?x \ ?y = ?y \ ?x"} \\ - \textbf{notes} & @{text lcomm}: @{text [quotes] "?x \ (?y \ ?z) = ?y \ (?x \ - ?z)"} -\end{array} \right]. -\end{align*% -} - Let $\pi_0(\F(l))$ denote the list of parameters defined in the - \textbf{fixes} elements of $\F(l)$ in the order of their occurrence. - The list of parameters of a locale expression $\pi(e)$ is defined as - follows: -\begin{align*% -} - \pi(l) & = \pi(\I(l)) \App \pi_0(\F(l)) \text{, for named locale $l$.} \\ - \pi(e\: q_1 \ldots q_n) & = \text{$[q_1, \ldots, q_n, p_{n+1}, \ldots, - p_{m}]$, where $\pi(e) = [p_1, \ldots, p_m]$.} \\ - \pi(e_1 + e_2) & = \pi(e_1) \App \pi(e_2) -\end{align*% -} - The operation $\App$ concatenates two lists but omits elements from - the second list that are also present in the first list. - It is not possible to rename more parameters than there - are present in an expression --- that is, $n \le m$ --- otherwise - the renaming is illegal. If $q_i - = \_$ then the $i$th entry of the resulting list is $p_i$. - - In the normalisation phase, imports of named locales are unfolded, and - renames and merges are recursively propagated to the imported locale - expressions. The result is a list of locale names, each with a full - list of parameters, where locale names occurring with the same parameter - list twice are removed. Let $\N$ denote normalisation. It is defined - by these equations: -\begin{align*% -} - \N(l) & = \N(\I(l)) \App [l\:\pi(l)] \text{, for named locale $l$.} \\ - \N(e\: q_1 \ldots q_n) & = \N(e)\:[q_1 \ldots q_n/\pi(e)] \\ - \N(e_1 + e_2) & = \N(e_1) \App \N(e_2) -\end{align*% -} - Normalisation yields a list of \emph{identifiers}. An identifier - consists of a locale name and a (possibly empty) list of parameters. - - In the second phase, the list of identifiers $\N(e)$ is converted to - a list of context elements $\C(e)$ by converting each identifier to - a list of context elements, and flattening the obtained list. - Conversion of the identifier $l\:q_1 \ldots q_n$ yields the list of - context elements $\F(l)$, but with the following modifications: -\begin{itemize} -\item - Rename the parameter in the $i$th \textbf{fixes} element of $\F(l)$ - to $q_i$, $i = 1, \ldots, n$. If the parameter name is actually - changed then delete the syntax annotation. Renaming a parameter may - also change its type. -\item - Perform the same renamings on all occurrences of parameters (fixed - variables) in \textbf{assumes}, \textbf{defines} and \textbf{notes} - elements. -\item - Qualify names of facts by $q_1\_\ldots\_q_n$. -\end{itemize} - The locale expression is \emph{well-formed} if it contains no - illegal renamings and the following conditions on $\C(e)$ hold, - otherwise the expression is rejected: -\begin{itemize} -\item Parameters in \textbf{fixes} are distinct; -\item Free variables in \textbf{assumes} and - \textbf{defines} occur in preceding \textbf{fixes};% -\footnote{This restriction is relaxed for contexts obtained with - \textbf{includes}, see Section~\ref{sec-includes}.} -\item Parameters defined in \textbf{defines} must neither occur in - preceding \textbf{assumes} nor \textbf{defines}. -\end{itemize} -*} - -subsection {* Examples *} - -text {* -\paragraph{Example 2.} - We obtain the context fragment $\C(@{text "comm_semi"})$ of the - locale @{text "comm_semi"}. First, the parameters are computed. -\begin{align*% -} - \pi(@{text "semi"}) & = [@{text prod}] \\ - \pi(@{text "comm_semi"}) & = \pi(@{text "semi"}) \App [] - = [@{text prod}] -\end{align*% -} - Next, the normal form of the locale expression - @{text "comm_semi"} is obtained. -\begin{align*% -} - \N(@{text "semi"}) & = [@{text semi} @{text prod}] \\ - \N(@{text "comm_semi"}) & = \N(@{text "semi"}) \App - [@{text "comm_semi prod"}] - = [@{text "semi prod"}, @{text "comm_semi prod"}] -\end{align*% -} - Converting this to a list of context elements leads to the - list~(\ref{eq-comm-semi}) shown in - Section~\ref{sec-locale-predicates}, but with fact names qualified - by @{text prod} --- for example, @{text "prod.assoc"}. - Qualification was omitted to keep the presentation simple. - Isabelle's scoping rules identify the most recent fact with - qualified name $x.a$ when a fact with name $a$ is requested. - -\paragraph{Example 3.} - The locale expression @{text "comm_semi sum"} involves - renaming. Computing parameters yields $\pi(@{text "comm_semi sum"}) - = [@{text sum}]$, the normal form is -\begin{align*% -} - \N(@{text "comm_semi sum"}) & = - \N(@{text "comm_semi"})[@{text sum}/@{text prod}] = - [@{text "semi sum"}, @{text "comm_semi sum"}] -\end{align*% -} - and the list of context elements -\[ -\begin{array}{ll} - \textbf{fixes} & @{text sum} :: @{typ [quotes] "['a, 'a] \ 'a"} - ~(\textbf{infixl}~@{text [quotes] "\"}~65) \\ - \textbf{assumes} & @{text [quotes] "semi sum"} \\ - \textbf{notes} & @{text sum.assoc}: @{text [quotes] "(?x \ ?y) \ ?z - = sum ?x (sum ?y ?z)"} \\ - \textbf{assumes} & @{text [quotes] "comm_semi_axioms sum"} \\ - \textbf{notes} & @{text sum.comm}: @{text [quotes] "?x \ ?y = - ?y \ ?x"} \\ - \textbf{notes} & @{text sum.lcomm}: @{text [quotes] "?x \ (?y \ ?z) - = ?y \ (?x \ ?z)"} -\end{array} -\] - -\paragraph{Example 4.} - The context defined by the locale @{text "semi_hom"} involves - merging two copies of @{text "comm_semi"}. We obtain the parameter - list and normal form: -\begin{align*% -} - \pi(@{text "semi_hom"}) & = \pi(@{text "comm_semi sum"} + - @{text "comm_semi"}) \App [@{text hom}] \\ - & = (\pi(@{text "comm_semi sum"}) \App \pi(@{text "comm_semi"})) - \App [@{text hom}] \\ - & = ([@{text sum}] \App [@{text prod}]) \App [@{text hom}] - = [@{text sum}, @{text prod}, @{text hom}] \\ - \N(@{text "semi_hom"}) & = - \N(@{text "comm_semi sum"} + @{text "comm_semi"}) \App \\ - & \phantom{==} - [@{text "semi_hom sum prod hom"}] \\ - & = (\N(@{text "comm_semi sum"}) \App \N(@{text "comm_semi"})) \App \\ - & \phantom{==} - [@{text "semi_hom sum prod hom"}] \\ - & = ([@{text "semi sum"}, @{text "comm_semi sum"}] \App %\\ -% & \phantom{==} - [@{text "semi prod"}, @{text "comm_semi prod"}]) \App \\ - & \phantom{==} - [@{text "semi_hom sum prod hom"}] \\ - & = [@{text "semi sum"}, @{text "comm_semi sum"}, - @{text "semi prod"}, @{text "comm_semi prod"}, \\ - & \phantom{==} - @{text "semi_hom sum prod hom"}]. -\end{align*% -} - Hence $\C(@{text "semi_hom"})$, shown below, is again well-formed. -\[ -\begin{array}{ll} - \textbf{fixes} & @{text sum} :: @{typ [quotes] "['a, 'a] \ 'a"} - ~(\textbf{infixl}~@{text [quotes] "\"}~65) \\ - \textbf{assumes} & @{text [quotes] "semi sum"} \\ - \textbf{notes} & @{text sum.assoc}: @{text [quotes] "(?x \ ?y) \ ?z - = ?x \ (?y \ ?z)"} \\ - \textbf{assumes} & @{text [quotes] "comm_semi_axioms sum"} \\ - \textbf{notes} & @{text sum.comm}: @{text [quotes] "?x \ ?y = ?y \ ?x"} \\ - \textbf{notes} & @{text sum.lcomm}: @{text [quotes] "?x \ (?y \ ?z) - = ?y \ (?x \ ?z)"} \\ - \textbf{fixes} & @{text prod} :: @{typ [quotes] "['b, 'b] \ 'b"} - ~(\textbf{infixl}~@{text [quotes] "\"}~70) \\ - \textbf{assumes} & @{text [quotes] "semi prod"} \\ - \textbf{notes} & @{text prod.assoc}: @{text [quotes] "?x \ ?y \ ?z = ?x \ (?y \ - ?z)"} \\ - \textbf{assumes} & @{text [quotes] "comm_semi_axioms prod"} \\ - \textbf{notes} & @{text prod.comm}: @{text [quotes] "?x \ ?y = ?y \ ?x"} \\ - \textbf{notes} & @{text prod.lcomm}: @{text [quotes] "?x \ (?y \ ?z) = ?y \ (?x \ - ?z)"} \\ - \textbf{fixes} & @{text hom} :: @{typ [quotes] "'a \ 'b"} \\ - \textbf{assumes} & @{text [quotes] "semi_hom_axioms sum"} \\ - \textbf{notes} & @{text "sum_prod_hom.hom"}: - @{text "hom (x \ y) = hom x \ hom y"} -\end{array} -\] - -\paragraph{Example 5.} - In this example, a locale expression leading to a list of context - elements that is not well-defined is encountered, and it is illustrated - how normalisation deals with multiple inheritance. - Consider the specification of monads (in the algebraic sense) - and monoids. -*} - -locale monad = - fixes prod :: "['a, 'a] \ 'a" (infixl "\" 70) and one :: 'a ("\" 100) - assumes l_one: "\ \ x = x" and r_one: "x \ \ = x" - -text {* - Monoids are both semigroups and monads and one would want to - specify them as locale expression @{text "semi + monad"}. - Unfortunately, this expression is not well-formed. Its normal form -\begin{align*% -} - \N(@{text "monad"}) & = [@{text "monad prod"}] \\ - \N(@{text "semi"}+@{text "monad"}) & = - \N(@{text "semi"}) \App \N(@{text "monad"}) - = [@{text "semi prod"}, @{text "monad prod"}] -\end{align*% -} - leads to a list containing the context element -\[ - \textbf{fixes}~@{text prod} :: @{typ [quotes] "['a, 'a] \ 'a"} - ~(\textbf{infixl}~@{text [quotes] "\"}~70) -\] - twice and thus violating the first criterion of well-formedness. To - avoid this problem, one can - introduce a new locale @{text magma} with the sole purpose of fixing the - parameter and defining its syntax. The specifications of semigroup - and monad are changed so that they import @{text magma}. -*} - -locale magma = fixes prod (infixl "\" 70) - -locale semi' = magma + assumes assoc: "(x \ y) \ z = x \ (y \ z)" -locale monad' = magma + fixes one ("\" 100) - assumes l_one: "\ \ x = x" and r_one: "x \ \ = x" - -text {* - Normalisation now yields -\begin{align*% -} - \N(@{text "semi' + monad'"}) & = - \N(@{text "semi'"}) \App \N(@{text "monad'"}) \\ - & = (\N(@{text magma}) \App [@{text "semi' prod"}]) \App - (\N(@{text magma}) \App [@{text "monad' prod"}]) \\ - & = [@{text "magma prod"}, @{text "semi' prod"}] \App - [@{text "magma prod"}, @{text "monad' prod"}]) \\ - & = [@{text "magma prod"}, @{text "semi' prod"}, - @{text "monad' prod"}] -\end{align*% -} - where the second occurrence of @{text "magma prod"} is eliminated. - The reader is encouraged to check, using the \textbf{print\_locale} - command, that the list of context elements generated from this is - indeed well-formed. - - It follows that importing - parameters is more flexible than fixing them using a context element. - The Locale package provides the predefined locale @{term var} that - can be used to import parameters if no - particular mixfix syntax is required. - Its definition is -\begin{center} - \textbf{locale} @{text var} = \textbf{fixes} @{text "x_"} -\end{center} - The use of the internal variable @{text "x_"} - enforces that the parameter is renamed before being used, because - internal variables may not occur in the input syntax. Using - @{text var}, the locale @{text magma} could have been replaced by - the locale expression -\begin{center} - @{text var} @{text prod} (\textbf{infixl} @{text [quotes] \} 70) -\end{center} - in the above locale declarations. -*} - -subsection {* Includes *} - -text {* -\label{sec-includes} - The context element \textbf{includes} takes a locale expression $e$ - as argument. It can only occur in long goals, where it - adds, like a target, locale context to the proof context. Unlike - with targets, the proved theorem is not stored - in the locale. Instead, it is exported immediately. -*} - -theorem lcomm2: - includes comm_semi shows "x \ (y \ z) = y \ (x \ z)" -proof - - have "x \ (y \ z) = (x \ y) \ z" by (simp add: assoc) - also have "\ = (y \ x) \ z" by (simp add: comm) - also have "\ = y \ (x \ z)" by (simp add: assoc) - finally show ?thesis . -qed - -text {* - This proof is identical to the proof of @{text lcomm}. The use of - \textbf{includes} provides the same context and facts as when using - @{term comm_semi} as target. On the other hand, @{thm [source] - lcomm2} is not added as a fact to the locale @{term comm_semi}, but - is directly visible in the theory. The theorem is -\[ - @{thm lcomm2}. -\] - Note that it is possible to - combine a target and (several) \textbf{includes} in a goal statement, thus - using contexts of several locales but storing the theorem in only - one of them. -*} - -section {* Structures *} - -text {* -\label{sec-structures} - The specifications of semigroups and monoids that served as examples - in previous sections modelled each operation of an algebraic - structure as a single parameter. This is rather inconvenient for - structures with many operations, and also unnatural. In accordance - to mathematical texts, one would rather fix two groups instead of - two sets of operations. - - The approach taken in Isabelle is to encode algebraic structures - with suitable types (in Isabelle/HOL usually records). An issue to - be addressed by - locales is syntax for algebraic structures. This is the purpose of - the \textbf{(structure)} annotation in \textbf{fixes}, introduced by - Wenzel. We illustrate this, independently of record types, with a - different formalisation of semigroups. - - Let @{text "'a semi_type"} be a not further specified type that - represents semigroups over the carrier type @{typ "'a"}. Let @{text - "s_op"} be an operation that maps an object of @{text "'a semi_type"} to - a binary operation. -*} - -typedecl 'a semi_type -consts s_op :: "['a semi_type, 'a, 'a] \ 'a" (infixl "\\" 70) - -text {* - Although @{text "s_op"} is a ternary operation, it is declared - infix. The syntax annotation contains the token @{text \} - (\verb.\.), which refers to the first argument. This syntax is only - effective in the context of a locale, and only if the first argument - is a - \emph{structural} parameter --- that is, a parameter with annotation - \textbf{(structure)}. The token has the effect of subscripting the - parameter --- by bracketing it between \verb.\<^bsub>. and \verb.\<^esub>.. - Additionally, the subscript of the first structural parameter may be - omitted, as in this specification of semigroups with structures: -*} - -locale comm_semi' = - fixes G :: "'a semi_type" (structure) - assumes assoc: "(x \ y) \ z = x \ (y \ z)" and comm: "x \ y = y \ x" - -text {* - Here @{text "x \ y"} is equivalent to @{text "x \\<^bsub>G\<^esub> y"} and - abbreviates @{text "s_op G x y"}. A specification of homomorphisms - requires a second structural parameter. -*} - -locale semi'_hom = comm_semi' + comm_semi' H + - fixes hom - assumes hom: "hom (x \ y) = hom x \\<^bsub>H\<^esub> hom y" - -text {* - The parameter @{text H} is defined in the second \textbf{fixes} - element of $\C(@{term "semi'_comm"})$. Hence @{text "\\<^bsub>H\<^esub>"} - abbreviates @{text "s_op H x y"}. The same construction can be done - with records instead of an \textit{ad-hoc} type. *} - -record 'a semi = prod :: "['a, 'a] \ 'a" (infixl "\\" 70) - -text {* - This declares the types @{typ "'a semi"} and @{typ "('a, 'b) - semi_scheme"}. The latter is an extensible record, where the second - type argument is the type of the extension field. For details on - records, see \cite{NipkowEtAl2002} Chapter~8.3. -*} - -locale semi_w_records = struct G + - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - -text {* - The type @{typ "('a, 'b) semi_scheme"} is inferred for the parameter @{term - G}. Using subtyping on records, the specification can be extended - to groups easily. -*} - -record 'a group = "'a semi" + - one :: "'a" ("l\" 100) - inv :: "'a \ 'a" ("inv\ _" [81] 80) -locale group_w_records = semi_w_records + - assumes l_one: "l \ x = x" and l_inv: "inv x \ x = l" - -text {* - Finally, the predefined locale -\begin{center} - \textbf{locale} \textit{struct} = \textbf{fixes} @{text "S_"} - \textbf{(structure)}. -\end{center} - is analogous to @{text "var"}. - More examples on the use of structures, including groups, rings and - polynomials can be found in the Isabelle distribution in the - session HOL-Algebra. -*} - -section {* Conclusions and Outlook *} - -text {* - Locales provide a simple means of modular reasoning. They enable to - abbreviate frequently occurring context statements and maintain facts - valid in these contexts. Importantly, using structures, they allow syntax to be - effective only in certain contexts, and thus to mimic common - practice in mathematics, where notation is chosen very flexibly. - This is also known as literate formalisation \cite{Bailey1998}. - Locale expressions allow to duplicate and merge - specifications. This is a necessity, for example, when reasoning about - homomorphisms. Normalisation makes it possible to deal with - diamond-shaped inheritance structures, and generally with directed - acyclic graphs. The combination of locales with record - types in higher-order logic provides an effective means for - specifying algebraic structures: locale import and record subtyping - provide independent hierarchies for specifications and structure - elements. Rich examples for this can be found in - the Isabelle distribution in the session HOL-Algebra. - - The primary reason for writing this report was to provide a better - understanding of locales in Isar. Wenzel provided hardly any - documentation, with the exception of \cite{Wenzel2002b}. The - present report should make it easier for users of Isabelle to take - advantage of locales. - - The report is also a base for future extensions. These include - improved syntax for structures. Identifying them by numbers seems - unnatural and can be confusing if more than two structures are - involved --- for example, when reasoning about universal - properties --- and numbering them by order of occurrence seems - arbitrary. Another desirable feature is \emph{instantiation}. One - may, in the course of a theory development, construct objects that - fulfil the specification of a locale. These objects are possibly - defined in the context of another locale. Instantiation should make it - simple to specialise abstract facts for the object under - consideration and to use the specified facts. - - A detailed comparison of locales with module systems in type theory - has not been undertaken yet, but could be beneficial. For example, - a module system for Coq has recently been presented by Chrzaszcz - \cite{Chrzaszcz2003,Chrzaszcz2004}. While the - latter usually constitute extensions of the calculus, locales are - a rather thin layer that does not change Isabelle's meta logic. - Locales mainly manage specifications and facts. Functors, like - the constructor for polynomial rings, remain objects of the - logic. - - \textbf{Acknowledgements.} Lawrence C.\ Paulson and Norbert - Schirmer made useful comments on a draft of this paper. -*} - -(*<*) -end -(*>*) \ No newline at end of file diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/ROOT.ML --- a/doc-src/Locales/Locales/ROOT.ML Tue Jun 03 11:55:35 2008 +0200 +++ b/doc-src/Locales/Locales/ROOT.ML Tue Jun 03 12:34:22 2008 +0200 @@ -1,6 +1,4 @@ -(* - no_document use_thy "ThisTheory"; - use_thy "ThatTheory"; -*) - -use_thy "Locales"; +no_document use_thy "GCD"; +use_thy "Examples1"; +use_thy "Examples2"; +use_thy "Examples3"; diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/Examples.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/document/Examples.tex Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,1403 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Examples}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Locales are based on contexts. A \emph{context} can be seen as a + formula schema +\[ + \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}{\isasymdots}x\isactrlsub n{\isachardot}\ {\isasymlbrakk}\ A\isactrlsub {\isadigit{1}}{\isacharsemicolon}\ {\isasymdots}\ {\isacharsemicolon}A\isactrlsub m\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymdots}} +\] + where variables \isa{x\isactrlsub {\isadigit{1}}}, \ldots, \isa{x\isactrlsub n} are called + \emph{parameters} and the premises $\isa{A\isactrlsub {\isadigit{1}}}, \ldots, + \isa{A\isactrlsub m}$ \emph{assumptions}. A formula \isa{C} + is a \emph{theorem} in the context if it is a conclusion +\[ +%\label{eq-fact-in-context} + \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}{\isasymdots}x\isactrlsub n{\isachardot}\ {\isasymlbrakk}\ A\isactrlsub {\isadigit{1}}{\isacharsemicolon}\ {\isasymdots}\ {\isacharsemicolon}A\isactrlsub m\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ C}. +\] + Isabelle/Isar's notion of context goes beyond this logical view. + Its contexts record, in a consecutive order, proved + conclusions along with attributes, which + may control proof procedures. Contexts also contain syntax information + for parameters and for terms depending on them.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Simple Locales% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Locales can be seen as persistent contexts. 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{\isacharunderscore}order}.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{locale}\isamarkupfalse% +\ partial{\isacharunderscore}order\ {\isacharequal}\isanewline +\ \ \ \ \isakeyword{fixes}\ le\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \ \ \isakeyword{assumes}\ refl\ {\isacharbrackleft}intro{\isacharcomma}\ simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ anti{\isacharunderscore}sym\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}\ x\ {\isasymsqsubseteq}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ x\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ trans\ {\isacharbrackleft}trans{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}\ x\ {\isasymsqsubseteq}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ z\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}% +\begin{isamarkuptext}% +The parameter of this locale is \isa{le}, with infix syntax + \isa{{\isasymsqsubseteq}}. There is an implicit type parameter \isa{{\isacharprime}a}. It + is not necessary to declare parameter types: most general types will + be inferred from the context elements for all parameters. + + The above declaration not only introduces the locale, it also + defines the \emph{locale predicate} \isa{partial{\isacharunderscore}order} with + definition \isa{partial{\isacharunderscore}order{\isacharunderscore}def}: + \begin{isabelle}% +\ \ partial{\isacharunderscore}order\ {\isacharquery}le\ {\isasymequiv}\isanewline +\isaindent{\ \ }{\isacharparenleft}{\isasymforall}x{\isachardot}\ {\isacharquery}le\ x\ x{\isacharparenright}\ {\isasymand}\isanewline +\isaindent{\ \ }{\isacharparenleft}{\isasymforall}x\ y{\isachardot}\ {\isacharquery}le\ x\ y\ {\isasymlongrightarrow}\ {\isacharquery}le\ y\ x\ {\isasymlongrightarrow}\ x\ {\isacharequal}\ y{\isacharparenright}\ {\isasymand}\isanewline +\isaindent{\ \ }{\isacharparenleft}{\isasymforall}x\ y\ z{\isachardot}\ {\isacharquery}le\ x\ y\ {\isasymlongrightarrow}\ {\isacharquery}le\ y\ z\ {\isasymlongrightarrow}\ {\isacharquery}le\ x\ z{\isacharparenright}% +\end{isabelle} + + 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. There are 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{fun}, \isakeyword{function} & recursive function \\ + \isakeyword{abbreviation} & syntactic abbreviation \\ + \isakeyword{theorem}, etc.\ & theorem statement with proof \\ + \isakeyword{theorems}, etc.\ & redeclaration of theorems +\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% +\ {\isacharparenleft}\isakeyword{in}\ partial{\isacharunderscore}order{\isacharparenright}\isanewline +\ \ \ \ less\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubset}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \ \ \isakeyword{where}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqsubset}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqsubseteq}\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +A definition in a locale depends on the locale parameters. + Here, a global constant \isa{partial{\isacharunderscore}order{\isachardot}less} is declared, which is lifted over the + locale parameter \isa{le}. Its definition is the global theorem + \isa{partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def}: + \begin{isabelle}% +\ \ partial{\isacharunderscore}order\ {\isacharquery}le\ {\isasymLongrightarrow}\isanewline +\isaindent{\ \ }partial{\isacharunderscore}order{\isachardot}less\ {\isacharquery}le\ {\isacharquery}x\ {\isacharquery}y\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}le\ {\isacharquery}x\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}% +\end{isabelle} + At the same time, the locale is extended by syntax information + hiding this construction in the context of the locale. That is, + \isa{partial{\isacharunderscore}order{\isachardot}less\ le} is printed and parsed as infix + \isa{{\isasymsqsubset}}. Finally, the conclusion of the definition is added to + the locale, \isa{less{\isacharunderscore}def}: + \begin{isabelle}% +\ \ {\isacharparenleft}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}% +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +As an example of a theorem statement in the locale, here is the + derivation of a transitivity law.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ partial{\isacharunderscore}order{\isacharparenright}\ less{\isacharunderscore}le{\isacharunderscore}trans\ {\isacharbrackleft}trans{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isasymlbrakk}\ x\ {\isasymsqsubset}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ z\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymsqsubset}\ z{\isachardoublequoteclose}\isanewline +% +\isadelimvisible +\ \ \ \ % +\endisadelimvisible +% +\isatagvisible +\isacommand{unfolding}\isamarkupfalse% +\ less{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}blast\ intro{\isacharcolon}\ trans{\isacharparenright}% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +In the context of the proof, assumptions and theorems of the + locale may be used. Attributes are effective: \isa{anti{\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% +% +\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. + + In the block below, notions of infimum and supremum together with + theorems are introduced for partial orders.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{context}\isamarkupfalse% +\ partial{\isacharunderscore}order\ \isakeyword{begin}\isanewline +\isanewline +\ \ \isacommand{definition}\isamarkupfalse% +\isanewline +\ \ \ \ is{\isacharunderscore}inf\ \isakeyword{where}\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i\ {\isacharequal}\isanewline +\ \ \ \ \ \ {\isacharparenleft}i\ {\isasymsqsubseteq}\ x\ {\isasymand}\ i\ {\isasymsqsubseteq}\ y\ {\isasymand}\ {\isacharparenleft}{\isasymforall}z{\isachardot}\ z\ {\isasymsqsubseteq}\ x\ {\isasymand}\ z\ {\isasymsqsubseteq}\ y\ {\isasymlongrightarrow}\ z\ {\isasymsqsubseteq}\ i{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \isacommand{definition}\isamarkupfalse% +\isanewline +\ \ \ \ is{\isacharunderscore}sup\ \isakeyword{where}\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s\ {\isacharequal}\isanewline +\ \ \ \ \ \ {\isacharparenleft}x\ {\isasymsqsubseteq}\ s\ {\isasymand}\ y\ {\isasymsqsubseteq}\ s\ {\isasymand}\ {\isacharparenleft}{\isasymforall}z{\isachardot}\ x\ {\isasymsqsubseteq}\ z\ {\isasymand}\ y\ {\isasymsqsubseteq}\ z\ {\isasymlongrightarrow}\ s\ {\isasymsqsubseteq}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadeliminvisible +\isanewline +\ \ % +\endisadeliminvisible +% +\isataginvisible +\isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}infI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ i\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\isanewline +\ \ \ \ \ \ {\isacharparenleft}{\isasymAnd}z{\isachardot}\ z\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ i{\isacharparenright}\ {\isasymLongrightarrow}\ is{\isacharunderscore}inf\ x\ y\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\ blast\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}inf{\isacharunderscore}lower\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i\ {\isasymLongrightarrow}\ {\isacharparenleft}i\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ i\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\ blast\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}inf{\isacharunderscore}greatest\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\ blast% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +\isanewline +% +\endisadeliminvisible +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ is{\isacharunderscore}inf{\isacharunderscore}uniq{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}is{\isacharunderscore}inf\ x\ y\ i{\isacharsemicolon}\ is{\isacharunderscore}inf\ x\ y\ i{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ i\ {\isacharequal}\ i{\isacharprime}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ inf{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ inf{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i{\isacharprime}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ anti{\isacharunderscore}sym{\isacharparenright}\isanewline +\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ inf{\isacharprime}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i\ {\isasymsqsubseteq}\ i{\isacharprime}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}greatest{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ inf\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ inf\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ inf\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i{\isacharprime}\ {\isasymsqsubseteq}\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}greatest{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ inf{\isacharprime}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i{\isacharprime}\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ inf{\isacharprime}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i{\isacharprime}\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\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{\isacharunderscore}inf{\isacharunderscore}related\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ is{\isacharunderscore}inf\ x\ y\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}supI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ s\ {\isasymLongrightarrow}\ y\ {\isasymsqsubseteq}\ s\ {\isasymLongrightarrow}\isanewline +\ \ \ \ \ \ {\isacharparenleft}{\isasymAnd}z{\isachardot}\ x\ {\isasymsqsubseteq}\ z\ {\isasymLongrightarrow}\ y\ {\isasymsqsubseteq}\ z\ {\isasymLongrightarrow}\ s\ {\isasymsqsubseteq}\ z{\isacharparenright}\ {\isasymLongrightarrow}\ is{\isacharunderscore}sup\ x\ y\ s{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\ blast\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}sup{\isacharunderscore}least\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s\ {\isasymLongrightarrow}\ x\ {\isasymsqsubseteq}\ z\ {\isasymLongrightarrow}\ y\ {\isasymsqsubseteq}\ z\ {\isasymLongrightarrow}\ s\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\ blast\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}sup{\isacharunderscore}upper\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s\ {\isasymLongrightarrow}\ {\isacharparenleft}x\ {\isasymsqsubseteq}\ s\ {\isasymLongrightarrow}\ y\ {\isasymsqsubseteq}\ s\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\ blast% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +\isanewline +% +\endisadeliminvisible +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ is{\isacharunderscore}sup{\isacharunderscore}uniq{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}is{\isacharunderscore}sup\ x\ y\ s{\isacharsemicolon}\ is{\isacharunderscore}sup\ x\ y\ s{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ s\ {\isacharequal}\ s{\isacharprime}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ sup{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ sup{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s{\isacharprime}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ anti{\isacharunderscore}sym{\isacharparenright}\isanewline +\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ sup\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}s\ {\isasymsqsubseteq}\ s{\isacharprime}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}least{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ sup{\isacharprime}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ s{\isacharprime}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ sup{\isacharprime}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ s{\isacharprime}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ sup{\isacharprime}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}s{\isacharprime}\ {\isasymsqsubseteq}\ s{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}least{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ sup\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ s{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ sup\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ s{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\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{\isacharunderscore}sup{\isacharunderscore}related\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ is{\isacharunderscore}sup\ x\ y\ y{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +\isanewline +% +\endisadeliminvisible +\isanewline +\ \ \isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +In fact, many more theorems need to be shown for a usable + theory of partial orders. The + above two serve as illustrative examples.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Two commands are provided to inspect locales: + \isakeyword{print\_locales} lists the names of all locales of the + theory; \isakeyword{print\_locale}~$n$ prints the parameters and + assumptions of locale $n$; \isakeyword{print\_locale!}~$n$ + additionally outputs the conclusions. + + The syntax of the locale commands discussed in this tutorial is + shown in Table~\ref{tab:commands}. See the + Isabelle/Isar Reference Manual~\cite{IsarRef} + for full documentation.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Import% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{sec:import} + + 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. + + With locales, this inheritance is achieved through \emph{import} of a + locale. The import comes before the context elements.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{locale}\isamarkupfalse% +\ lattice\ {\isacharequal}\ partial{\isacharunderscore}order\ {\isacharplus}\isanewline +\ \ \ \ \isakeyword{assumes}\ ex{\isacharunderscore}inf{\isacharcolon}\ {\isachardoublequoteopen}{\isasymexists}inf{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}inf\ le\ x\ y\ inf{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ ex{\isacharunderscore}sup{\isacharcolon}\ {\isachardoublequoteopen}{\isasymexists}sup{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}sup\ le\ x\ y\ sup{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{begin}% +\begin{isamarkuptext}% +Note that the assumptions above refer to the predicates for infimum + and supremum defined in \isa{partial{\isacharunderscore}order}. In the current + implementation of locales, syntax from definitions of the imported + locale is unavailable in the locale declaration, neither are their + names. Hence we refer to the constants of the theory. The names + and syntax is available below, in the context of the locale.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{definition}\isamarkupfalse% +\isanewline +\ \ \ \ meet\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isacharequal}\ {\isacharparenleft}THE\ inf{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ inf{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \isacommand{definition}\isamarkupfalse% +\isanewline +\ \ \ \ join\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{5}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isacharequal}\ {\isacharparenleft}THE\ sup{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ sup{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadeliminvisible +\isanewline +\ \ % +\endisadeliminvisible +% +\isataginvisible +\isacommand{lemma}\isamarkupfalse% +\ meet{\isacharunderscore}equality\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i\ {\isasymLongrightarrow}\ x\ {\isasymsqinter}\ y\ {\isacharequal}\ i{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}unfold\ meet{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}THE\ i{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ i{\isacharparenright}\ {\isacharequal}\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}uniq\ {\isacharbrackleft}OF\ {\isacharunderscore}\ {\isacharbackquoteopen}is{\isacharunderscore}inf\ x\ y\ i{\isacharbackquoteclose}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ meetI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}i\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ i\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}z{\isachardot}\ z\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ i{\isacharparenright}\ {\isasymLongrightarrow}\ x\ {\isasymsqinter}\ y\ {\isacharequal}\ i{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ meet{\isacharunderscore}equality{\isacharcomma}\ rule\ is{\isacharunderscore}infI{\isacharparenright}\ blast{\isacharplus}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}inf{\isacharunderscore}meet\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ {\isacharparenleft}x\ {\isasymsqinter}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}unfold\ meet{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ ex{\isacharunderscore}inf\ \isacommand{obtain}\isamarkupfalse% +\ i\ \isakeyword{where}\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ i{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ {\isacharparenleft}THE\ i{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ theI{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}uniq\ {\isacharbrackleft}OF\ {\isacharunderscore}\ {\isacharbackquoteopen}is{\isacharunderscore}inf\ x\ y\ i{\isacharbackquoteclose}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ meet{\isacharunderscore}left\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}lower{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}meet{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ meet{\isacharunderscore}right\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}lower{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}meet{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ meet{\isacharunderscore}le\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isasymlbrakk}\ z\ {\isasymsqsubseteq}\ x{\isacharsemicolon}\ z\ {\isasymsqsubseteq}\ y\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ z\ {\isasymsqsubseteq}\ x\ {\isasymsqinter}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}greatest{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}inf{\isacharunderscore}meet{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ join{\isacharunderscore}equality\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s\ {\isasymLongrightarrow}\ x\ {\isasymsqunion}\ y\ {\isacharequal}\ s{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}unfold\ join{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}THE\ s{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ s{\isacharparenright}\ {\isacharequal}\ s{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}uniq\ {\isacharbrackleft}OF\ {\isacharunderscore}\ {\isacharbackquoteopen}is{\isacharunderscore}sup\ x\ y\ s{\isacharbackquoteclose}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ joinI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ s\ {\isasymLongrightarrow}\ y\ {\isasymsqsubseteq}\ s\ {\isasymLongrightarrow}\isanewline +\ \ \ \ \ \ {\isacharparenleft}{\isasymAnd}z{\isachardot}\ x\ {\isasymsqsubseteq}\ z\ {\isasymLongrightarrow}\ y\ {\isasymsqsubseteq}\ z\ {\isasymLongrightarrow}\ s\ {\isasymsqsubseteq}\ z{\isacharparenright}\ {\isasymLongrightarrow}\ x\ {\isasymsqunion}\ y\ {\isacharequal}\ s{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ join{\isacharunderscore}equality{\isacharcomma}\ rule\ is{\isacharunderscore}supI{\isacharparenright}\ blast{\isacharplus}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ is{\isacharunderscore}sup{\isacharunderscore}join\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}unfold\ join{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ ex{\isacharunderscore}sup\ \isacommand{obtain}\isamarkupfalse% +\ s\ \isakeyword{where}\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ s{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ {\isacharparenleft}THE\ s{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ s{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ theI{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}uniq\ {\isacharbrackleft}OF\ {\isacharunderscore}\ {\isacharbackquoteopen}is{\isacharunderscore}sup\ x\ y\ s{\isacharbackquoteclose}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ join{\isacharunderscore}left\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}upper{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}join{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ join{\isacharunderscore}right\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}upper{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}join{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ join{\isacharunderscore}le\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isasymlbrakk}\ x\ {\isasymsqsubseteq}\ z{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ z\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymsqunion}\ y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}least{\isacharparenright}\ {\isacharparenleft}rule\ is{\isacharunderscore}sup{\isacharunderscore}join{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqinter}\ y{\isacharparenright}\ {\isasymsqinter}\ z\ {\isacharequal}\ x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ meetI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ x\ {\isasymsqinter}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ y\ {\isasymsqinter}\ z{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ y\ {\isasymsqinter}\ z{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ w\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ x\ {\isasymsqinter}\ y{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ x\ {\isasymsqinter}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ y\ {\isasymsqinter}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ x\ {\isasymsqinter}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}w\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}commute{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isacharequal}\ y\ {\isasymsqinter}\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ meetI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqinter}\ x\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqinter}\ x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ y\ {\isasymsqinter}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}join{\isacharunderscore}absorb{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ meetI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isasymsqunion}\ z\ {\isacharequal}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ joinI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ y\ {\isasymsqunion}\ z{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ y\ {\isasymsqunion}\ z{\isachardoublequoteclose}\ \ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ w\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqunion}\ z\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ w{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}commute{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isacharequal}\ y\ {\isasymsqunion}\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ joinI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymsqunion}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ y\ {\isasymsqunion}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqunion}\ x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}meet{\isacharunderscore}absorb{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ {\isacharparenleft}x\ {\isasymsqinter}\ y{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ joinI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}idem{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}x\ {\isasymsqunion}\ {\isacharparenleft}x\ {\isasymsqinter}\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ meet{\isacharunderscore}join{\isacharunderscore}absorb{\isacharparenright}\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ {\isacharparenleft}x\ {\isasymsqinter}\ x{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ join{\isacharunderscore}meet{\isacharunderscore}absorb{\isacharparenright}\isanewline +\ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}related\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ x\ {\isasymsqinter}\ y\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ meetI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}related{\isadigit{2}}\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ x\ {\isasymsqinter}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}drule\ meet{\isacharunderscore}related{\isacharparenright}\ {\isacharparenleft}simp\ add{\isacharcolon}\ meet{\isacharunderscore}commute{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}related\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ x\ {\isasymsqunion}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ joinI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ z\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}related{\isadigit{2}}\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ x\ {\isasymLongrightarrow}\ x\ {\isasymsqunion}\ y\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}drule\ join{\isacharunderscore}related{\isacharparenright}\ {\isacharparenleft}simp\ add{\isacharcolon}\ join{\isacharunderscore}commute{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}connection{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqsubseteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqinter}\ y\ {\isacharequal}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}connection{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqsubseteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y\ {\isacharequal}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ x\ {\isasymsqunion}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ meet{\isacharunderscore}connection{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqsubseteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y\ {\isasymsqinter}\ x\ {\isacharequal}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ meet{\isacharunderscore}commute\ meet{\isacharunderscore}connection\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isanewline +\ \ \isacommand{theorem}\isamarkupfalse% +\ join{\isacharunderscore}connection{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqsubseteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y\ {\isacharequal}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ join{\isacharunderscore}commute\ join{\isacharunderscore}connection\ \isacommand{by}\isamarkupfalse% +\ simp% +\begin{isamarkuptext}% +Naming according to Jacobson I, p.\ 459.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemmas}\isamarkupfalse% +\ L{\isadigit{1}}\ {\isacharequal}\ join{\isacharunderscore}commute\ meet{\isacharunderscore}commute\isanewline +\ \ \isacommand{lemmas}\isamarkupfalse% +\ L{\isadigit{2}}\ {\isacharequal}\ join{\isacharunderscore}assoc\ meet{\isacharunderscore}assoc\isanewline +\ \ \isanewline +\ \ \isacommand{lemmas}\isamarkupfalse% +\ L{\isadigit{4}}\ {\isacharequal}\ join{\isacharunderscore}meet{\isacharunderscore}absorb\ meet{\isacharunderscore}join{\isacharunderscore}absorb% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +\isanewline +% +\endisadeliminvisible +\isanewline +\ \ \isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Locales for total orders and distributive lattices follow. + Each comes with an example theorem.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{locale}\isamarkupfalse% +\ total{\isacharunderscore}order\ {\isacharequal}\ partial{\isacharunderscore}order\ {\isacharplus}\isanewline +\ \ \ \ \isakeyword{assumes}\ total{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymor}\ y\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ total{\isacharunderscore}order{\isacharparenright}\ less{\isacharunderscore}total{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubset}\ y\ {\isasymor}\ x\ {\isacharequal}\ y\ {\isasymor}\ y\ {\isasymsqsubset}\ x{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ total\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}unfold\ less{\isacharunderscore}def{\isacharparenright}\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\ \ \isacommand{locale}\isamarkupfalse% +\ distrib{\isacharunderscore}lattice\ {\isacharequal}\ lattice\ {\isacharplus}\isanewline +\ \ \ \ \isakeyword{assumes}\ meet{\isacharunderscore}distr{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}lattice{\isachardot}meet\ le\ x\ {\isacharparenleft}lattice{\isachardot}join\ le\ y\ z{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ lattice{\isachardot}join\ le\ {\isacharparenleft}lattice{\isachardot}meet\ le\ x\ y{\isacharparenright}\ {\isacharparenleft}lattice{\isachardot}meet\ le\ x\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ distrib{\isacharunderscore}lattice{\isacharparenright}\ join{\isacharunderscore}distr{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isasymsqinter}\ {\isacharparenleft}x\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\ \ \isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ {\isacharparenleft}x\ {\isasymsqinter}\ z{\isacharparenright}{\isacharparenright}\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ L{\isadigit{4}}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x\ {\isasymsqunion}\ {\isacharparenleft}{\isacharparenleft}x\ {\isasymsqinter}\ z{\isacharparenright}\ {\isasymsqunion}\ {\isacharparenleft}y\ {\isasymsqinter}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ L{\isadigit{2}}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x\ {\isasymsqunion}\ {\isacharparenleft}{\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isasymsqinter}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ L{\isadigit{1}}\ meet{\isacharunderscore}distr{\isacharparenright}\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}{\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isasymsqinter}\ x{\isacharparenright}\ {\isasymsqunion}\ {\isacharparenleft}{\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isasymsqinter}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ L{\isadigit{1}}\ L{\isadigit{4}}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isasymsqinter}\ {\isacharparenleft}x\ {\isasymsqunion}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ meet{\isacharunderscore}distr{\isacharparenright}\isanewline +\ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +The locale hierachy obtained through these declarations is shown in Figure~\ref{fig:lattices}(a). + +\begin{figure} +\hrule \vspace{2ex} +\begin{center} +\subfigure[Declared hierachy]{ +\begin{tikzpicture} + \node (po) at (0,0) {\isa{partial{\isacharunderscore}order}}; + \node (lat) at (-1.5,-1) {\isa{lattice}}; + \node (dlat) at (-1.5,-2) {\isa{distrib{\isacharunderscore}lattice}}; + \node (to) at (1.5,-1) {\isa{total{\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{\isacharunderscore}order}}; + \node (lat) at (0,-1) {\isa{lattice}}; + \node (dlat) at (-1.5,-2) {\isa{distrib{\isacharunderscore}lattice}}; + \node (to) at (1.5,-2) {\isa{total{\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{\isacharunderscore}order}}; + \node (lat) at (0,-1) {\isa{lattice}}; + \node (dlat) at (0,-2) {\isa{distrib{\isacharunderscore}lattice}}; + \node (to) at (0,-3) {\isa{total{\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% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{sec:changing-the-hierarchy} + + Total orders are lattices. Hence, by deriving the lattice + axioms for total orders, the hierarchy may be changed + and \isa{lattice} be placed between \isa{partial{\isacharunderscore}order} + and \isa{total{\isacharunderscore}order}, as shown in Figure~\ref{fig:lattices}(b). + Changes to the locale hierarchy may be declared + with the \isakeyword{interpretation} command.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +\ \ % +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ total{\isacharunderscore}order\ {\isasymsubseteq}\ lattice% +\begin{isamarkuptxt}% +This enters the context of locale \isa{total{\isacharunderscore}order}, in + which the goal \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ lattice\ op\ {\isasymsqsubseteq}% +\end{isabelle} must be shown. First, the + locale predicate needs to be unfolded --- for example using its + definition or by introduction rules + provided by the locale package. The methods \isa{intro{\isacharunderscore}locales} + and \isa{unfold{\isacharunderscore}locales} automate this. They are aware of the + current context and dependencies between locales and automatically + discharge goals implied by these. While \isa{unfold{\isacharunderscore}locales} + always unfolds locale predicates to assumptions, \isa{intro{\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}, hence \isa{unfold{\isacharunderscore}locales} + is appropriate.% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{proof}\isamarkupfalse% +\ unfold{\isacharunderscore}locales% +\begin{isamarkuptxt}% +Since both \isa{lattice} and \isa{total{\isacharunderscore}order} + inherit \isa{partial{\isacharunderscore}order}, the assumptions of the latter are + discharged, and the only subgoals that remain are the assumptions + introduced in \isa{lattice} \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}inf{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ inf\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}sup{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ sup% +\end{isabelle} + The proof for the first subgoal is% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ y\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ total\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}is{\isacharunderscore}inf\ x\ y\ {\isacharparenleft}if\ x\ {\isasymsqsubseteq}\ y\ then\ x\ else\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymexists}inf{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ inf{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\begin{isamarkuptxt}% +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% +\ {\isachardoublequoteopen}is{\isacharunderscore}sup\ x\ y\ {\isacharparenleft}if\ x\ {\isasymsqsubseteq}\ y\ then\ y\ else\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymexists}sup{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ sup{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimvisible +\ % +\endisadelimvisible +% +\isatagvisible +\isacommand{qed}\isamarkupfalse% +% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +Similarly, total orders are distributive lattices.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{interpretation}\isamarkupfalse% +\ total{\isacharunderscore}order\ {\isasymsubseteq}\ distrib{\isacharunderscore}lattice\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ y\ z\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ {\isacharparenleft}y\ {\isasymsqunion}\ z{\isacharparenright}\ {\isacharequal}\ x\ {\isasymsqinter}\ y\ {\isasymsqunion}\ x\ {\isasymsqinter}\ z{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}l\ {\isacharequal}\ {\isacharquery}r{\isachardoublequoteclose}{\isacharparenright}% +\begin{isamarkuptxt}% +Jacobson I, p.\ 462% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ c{\isacharcolon}\ {\isachardoublequoteopen}y\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\ {\isachardoublequoteopen}z\ {\isasymsqsubseteq}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ c\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharquery}l\ {\isacharequal}\ y\ {\isasymsqunion}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ c\ join{\isacharunderscore}connection{\isadigit{2}}\ join{\isacharunderscore}related{\isadigit{2}}\ meet{\isacharunderscore}related{\isadigit{2}}\ total{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{from}\isamarkupfalse% +\ c\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharquery}r{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ meet{\isacharunderscore}related{\isadigit{2}}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharquery}l\ {\isacharequal}\ {\isacharquery}r{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\ \isacommand{{\isacharbraceright}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ c{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymor}\ x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ c\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharquery}l\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ join{\isacharunderscore}connection{\isadigit{2}}\ join{\isacharunderscore}related{\isadigit{2}}\ meet{\isacharunderscore}connection\ total\ trans{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{from}\isamarkupfalse% +\ c\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharquery}r{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ join{\isacharunderscore}commute\ join{\isacharunderscore}related{\isadigit{2}}\ meet{\isacharunderscore}connection\ meet{\isacharunderscore}related{\isadigit{2}}\ total{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharquery}l\ {\isacharequal}\ {\isacharquery}r{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\ \isacommand{{\isacharbraceright}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse% +\ \isacommand{note}\isamarkupfalse% +\ total\isanewline +\ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\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% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/Examples1.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/document/Examples1.tex Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,150 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Examples{\isadigit{1}}}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Use of Locales in Theories and Proofs% +} +\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}. + + The changes of the locale + hierarchy from the previous sections are examples of + interpretations. The command \isakeyword{interpretation} $l_1 + \subseteq l_2$ is said to \emph{interpret} locale $l_2$ in the + context of $l_1$. It causes all theorems of $l_2$ to be made + available in $l_1$. The interpretation is \emph{dynamic}: not only + theorems already present in $l_2$ are available in $l_1$. Theorems + that will be added to $l_2$ in future will automatically be + propagated to $l_1$. + + Locales can also be interpreted in the contexts of theories and + structured proofs. These interpretations are dynamic, too. + Theorems added to locales will be propagated to theories. + In this section the interpretation in + theories is illustrated; interpretation in proofs is analogous. + As an example consider, the type of natural numbers \isa{nat}. The + order relation \isa{{\isasymle}} is a total order over \isa{nat}, + divisibility \isa{dvd} is a distributive lattice. + + We start with the + interpretation that \isa{{\isasymle}} is a partial order. The facilities of + the interpretation command are explored in three versions.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{First Version: Replacement of Parameters Only% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{sec:po-first} + + In the most basic form, interpretation just replaces the locale + parameters by terms. The following command interprets the locale + \isa{partial{\isacharunderscore}order} in the global context of the theory. The + parameter \isa{le} is replaced by \isa{op\ {\isasymle}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +\ \ % +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}% +\begin{isamarkuptxt}% +The locale name is succeeded by a \emph{parameter + instantiation}. In general, this is a list of terms, which refer to + the parameters in the order of declaration in the locale. The + locale name is preceded by an optional \emph{interpretation prefix}, + which is used to qualify names from the locale in the global + context. + + The command creates the goal \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ partial{\isacharunderscore}order\ op\ {\isasymle}% +\end{isabelle} which can be shown + easily:% +\footnote{Note that \isa{op} binds tighter than functions + application: parentheses around \isa{op\ {\isasymle}} are not necessary.}% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\ auto% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +Now theorems from the locale are available in the theory, + interpreted for natural numbers, for example \isa{nat{\isachardot}trans}: \begin{isabelle}% +\ \ {\isasymlbrakk}{\isacharquery}x\ {\isasymle}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymle}\ {\isacharquery}z% +\end{isabelle} + + In order to avoid unwanted hiding of theorems, interpretation + accepts a qualifier, \isa{nat} in the example, which is applied to + all names processed by the + interpretation. If present the qualifier is mandatory --- that is, + the above theorem cannot be referred to simply as \isa{trans}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Second Version: Replacement of Definitions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The above interpretation also creates the theorem + \isa{nat{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}% +\ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}z{\isasymrbrakk}\isanewline +\isaindent{\ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}z% +\end{isabelle} + Here, \isa{partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}} + represents the strict order, although \isa{{\isacharless}} is defined on + \isa{nat}. Interpretation enables to map concepts + introduced in locales through definitions to the corresponding + concepts of the theory.% +\footnote{This applies not only to \isakeyword{definition} but also to + \isakeyword{inductive}, \isakeyword{fun} and \isakeyword{function}.}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/Examples2.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/document/Examples2.tex Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,90 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Examples{\isadigit{2}}}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\begin{isamarkuptext}% +This is achieved by unfolding suitable equations during + interpretation. These equations are given after the keyword + \isakeyword{where} and require proofs. The revised command, + replacing \isa{{\isasymsqsubset}} by \isa{{\isacharless}}, is:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +% +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}% +\begin{isamarkuptxt}% +The goals are \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ partial{\isacharunderscore}order\ op\ {\isasymle}\isanewline +\ {\isadigit{2}}{\isachardot}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}% +\end{isabelle} + The proof that \isa{{\isasymle}} is a partial order is a above.% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\ auto% +\begin{isamarkuptxt}% +The second goal is shown by unfolding the + definition of \isa{partial{\isacharunderscore}order{\isachardot}less}.% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def\ {\isacharbrackleft}OF\ {\isacharbackquoteopen}partial{\isacharunderscore}order\ op\ {\isasymle}{\isacharbackquoteclose}{\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 a locale. + Hence, the correct interpretation of \isa{partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def} is obtained manually with \isa{OF}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/Examples3.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Locales/Locales/document/Examples3.tex Tue Jun 03 12:34:22 2008 +0200 @@ -0,0 +1,850 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Examples{\isadigit{3}}}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsubsection{Third Version: Local Interpretation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In the above example, the fact that \isa{{\isasymle}} is a partial + order for the natural numbers was used in the proof of the + second goal. In general, proofs of the equations may involve + theorems implied by the fact the assumptions of the instantiated + locale hold for the instantiating structure. If these theorems have + been shown abstractly in the locale they can be made available + conveniently in the context through an auxiliary local interpretation (keyword + \isakeyword{interpret}). This interpretation is inside the proof of the global + interpretation. The third revision of the example illustrates this.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +% +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ {\isacharparenleft}op\ {\isasymle}{\isacharparenright}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\ auto\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{interpret}\isamarkupfalse% +\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ {\isacharparenleft}op\ {\isasymle}{\isacharparenright}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ nat{\isachardot}less{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +The inner interpretation does not require an + elaborate new proof, it is immediate from the preceeding fact and + proved with ``.''. + This interpretation enriches the local proof context by + the very theorems also obtained in the interpretation from + Section~\ref{sec:po-first}, and \isa{nat{\isachardot}less{\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 closing \isakeyword{qed} --- and are + then replaced by those with the desired substitutions of the strict + order.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Further Interpretations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Further interpretations are necessary to reuse theorems from + the other locales. In \isa{lattice} the operations \isa{{\isasymsqinter}} and + \isa{{\isasymsqunion}} are substituted by \isa{min} and + \isa{max}. The entire proof for the + interpretation is reproduced in order to give an example of a more + elaborate interpretation proof.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +% +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptxt}% +We have already shown that this is a partial order,% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ unfold{\isacharunderscore}locales% +\begin{isamarkuptxt}% +hence only the lattice axioms remain to be shown: \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}inf{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}inf\ op\ {\isasymle}\ x\ y\ inf\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}sup{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}sup\ op\ {\isasymle}\ x\ y\ sup% +\end{isabelle} After unfolding \isa{is{\isacharunderscore}inf} and \isa{is{\isacharunderscore}sup},% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def\ nat{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}% +\begin{isamarkuptxt}% +the goals become \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}inf{\isasymle}x{\isachardot}\ inf\ {\isasymle}\ y\ {\isasymand}\ {\isacharparenleft}{\isasymforall}z{\isachardot}\ z\ {\isasymle}\ x\ {\isasymand}\ z\ {\isasymle}\ y\ {\isasymlongrightarrow}\ z\ {\isasymle}\ inf{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}sup{\isasymge}x{\isachardot}\ y\ {\isasymle}\ sup\ {\isasymand}\ {\isacharparenleft}{\isasymforall}z{\isachardot}\ x\ {\isasymle}\ z\ {\isasymand}\ y\ {\isasymle}\ z\ {\isasymlongrightarrow}\ sup\ {\isasymle}\ z{\isacharparenright}% +\end{isabelle} which can be solved + by Presburger arithmetic.% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ arith{\isacharplus}% +\begin{isamarkuptxt}% +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% +\ nat{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}bestsimp\ simp{\isacharcolon}\ nat{\isachardot}meet{\isacharunderscore}def\ nat{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}bestsimp\ simp{\isacharcolon}\ nat{\isachardot}join{\isacharunderscore}def\ nat{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +That the relation \isa{{\isasymle}} is a total order completes this + sequence of interpretations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +% +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharcolon}\ total{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ unfold{\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:nat-lattice}. + +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{l} + \isa{nat{\isachardot}less{\isacharunderscore}def} from locale \isa{partial{\isacharunderscore}order}: \\ + \quad \isa{{\isacharparenleft}{\isacharquery}x\ {\isacharless}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ {\isasymle}\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}} \\ + \isa{nat{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\ + \quad \isa{min\ {\isacharquery}x\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}x} \\ + \isa{nat{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\ + \quad \isa{max\ {\isacharquery}x\ {\isacharparenleft}min\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ min\ {\isacharparenleft}max\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}max\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\ + \isa{nat{\isachardot}less{\isacharunderscore}total} from locale \isa{total{\isacharunderscore}order}: \\ + \quad \isa{{\isacharquery}x\ {\isacharless}\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}y\ {\isacharless}\ {\isacharquery}x} +\end{tabular} +\end{center} +\hrule +\caption{Interpreted theorems for \isa{{\isasymle}} on the natural numbers.} +\label{tab:nat-lattice} +\end{table} + + Note that since the locale hierarchy reflects that total orders are + distributive lattices, an explicit interpretation of distributive + lattices for the order relation on natural numbers is not neccessary. + + Why not push this idea further and just give the last interpretation + as a single interpretation instead of the sequence of three? The + reasons for this are twofold: +\begin{itemize} +\item + Often it is easier to work in an incremental fashion, because later + interpretations require theorems provided by earlier + interpretations. +\item + Assume that a definition is made in some locale $l_1$, and that $l_2$ + imports $l_1$. Let an equation for the definition be + proved in an interpretation of $l_2$. The equation will be unfolded + in interpretations of theorems added to $l_2$ or below in the import + hierarchy, but not for theorems added above $l_2$. + Hence, an equation interpreting a definition should always be given in + an interpretation of the locale where the definition is made, not in + an interpretation of a locale further down the hierarchy. +\end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Lattice \isa{dvd} on \isa{nat}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Divisibility on the natural numbers is a distributive lattice + but not a total order. Interpretation again proceeds + incrementally.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\ {\isacharparenleft}auto\ simp{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{interpret}\isamarkupfalse% +\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}less{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ auto\isanewline +\ \ \ \ \isacommand{done}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Note that there is no symbol for strict divisibility. Instead, + interpretation substitutes \isa{x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y}.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ gcd\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ lcm\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}gcd\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ auto\ {\isacharbrackleft}{\isadigit{1}}{\isacharbrackright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}lcm\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}dvd{\isadigit{1}}\ lcm{\isacharunderscore}dvd{\isadigit{2}}\ lcm{\isacharunderscore}least{\isacharparenright}\isanewline +\ \ \ \ \isacommand{done}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{interpret}\isamarkupfalse% +\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ gcd\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}meet{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ lcm\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}join{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}dvd{\isadigit{1}}\ lcm{\isacharunderscore}dvd{\isadigit{2}}\ lcm{\isacharunderscore}least{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Equations \isa{nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq} and \isa{nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq} are named since they are handy in the proof of + the subsequent interpretation.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\ set\ quick{\isacharunderscore}and{\isacharunderscore}dirty\ {\isacharverbatimclose}\isanewline +\isanewline +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}gcd\ {\isacharparenleft}x{\isacharcomma}\ lcm\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ lcm\ {\isacharparenleft}gcd\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isacharcomma}\ gcd\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{sorry}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\ reset\ quick{\isacharunderscore}and{\isacharunderscore}dirty\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +\isanewline +% +\isadelimvisible +\ \ \isanewline +% +\endisadelimvisible +% +\isatagvisible +\isacommand{interpretation}\isamarkupfalse% +\ nat{\isacharunderscore}dvd{\isacharcolon}\isanewline +\ \ distrib{\isacharunderscore}lattice\ {\isacharbrackleft}{\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}{\isacharbrackright}\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ unfold{\isacharunderscore}locales% +\begin{isamarkuptxt}% +\begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y\ z{\isachardot}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }lattice{\isachardot}meet\ op\ dvd\ x\ {\isacharparenleft}lattice{\isachardot}join\ op\ dvd\ y\ z{\isacharparenright}\ {\isacharequal}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }lattice{\isachardot}join\ op\ dvd\ {\isacharparenleft}lattice{\isachardot}meet\ op\ dvd\ x\ y{\isacharparenright}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }{\isacharparenleft}lattice{\isachardot}meet\ op\ dvd\ x\ z{\isacharparenright}% +\end{isabelle}% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}% +\begin{isamarkuptxt}% +\begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y\ z{\isachardot}\ gcd\ {\isacharparenleft}x{\isacharcomma}\ lcm\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ lcm\ {\isacharparenleft}gcd\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isacharcomma}\ gcd\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}{\isacharparenright}% +\end{isabelle}% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharparenright}\ \isacommand{done}\isamarkupfalse% +% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +Theorems that are available in the theory after these + interpretations are shown in Table~\ref{tab:nat-dvd-lattice}. + +\begin{table} +\hrule +\vspace{2ex} +\begin{center} +\begin{tabular}{l} + \isa{nat{\isacharunderscore}dvd{\isachardot}less{\isacharunderscore}def} from locale \isa{partial{\isacharunderscore}order}: \\ + \quad \isa{{\isacharparenleft}{\isacharquery}x\ dvd\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ dvd\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}} \\ + \isa{nat{\isacharunderscore}dvd{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\ + \quad \isa{gcd\ {\isacharparenleft}{\isacharquery}x{\isacharcomma}\ {\isacharquery}y{\isacharparenright}\ dvd\ {\isacharquery}x} \\ + \isa{nat{\isacharunderscore}dvd{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\ + \quad \isa{lcm\ {\isacharparenleft}{\isacharquery}x{\isacharcomma}\ gcd\ {\isacharparenleft}{\isacharquery}y{\isacharcomma}\ {\isacharquery}z{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}lcm\ {\isacharparenleft}{\isacharquery}x{\isacharcomma}\ {\isacharquery}y{\isacharparenright}{\isacharcomma}\ lcm\ {\isacharparenleft}{\isacharquery}x{\isacharcomma}\ {\isacharquery}z{\isacharparenright}{\isacharparenright}} \\ +\end{tabular} +\end{center} +\hrule +\caption{Interpreted theorems for \isa{dvd} on the natural numbers.} +\label{tab:nat-dvd-lattice} +\end{table}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +The full syntax of the interpretation commands is shown in + Table~\ref{tab:commands}. The grammar refers to + \textit{expr}, which stands for a \emph{locale} expression. Locale + expressions are discussed in Section~\ref{sec:expressions}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Locale Expressions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{sec:expressions} + + A map \isa{{\isasymphi}} between partial orders \isa{{\isasymsqsubseteq}} and \isa{{\isasympreceq}} + is called order preserving if \isa{x\ {\isasymsqsubseteq}\ y} implies \isa{{\isasymphi}\ x\ {\isasympreceq}\ {\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. + + Inspecting the grammar of locale commands in + Table~\ref{tab:commands} reveals that the import of a locale can be + more than just a single locale. In general, the import is a + \emph{locale expression}. Locale expressions enable to combine locales + and rename parameters. A locale name is a locale expression. If + $e_1$ and $e_2$ are locale expressions then $e_1 + e_2$ is their + \emph{merge}. If $e$ is an expression, then $e\:q_1 \ldots q_n$ is + a \emph{renamed expression} where the parameters in $e$ are renamed + to $q_1 \ldots q_n$. Using a locale expression, a locale for order + preserving maps can be declared in the following way.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{locale}\isamarkupfalse% +\ order{\isacharunderscore}preserving\ {\isacharequal}\isanewline +\ \ \ \ partial{\isacharunderscore}order\ {\isacharplus}\ partial{\isacharunderscore}order\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline +\ \ \ \ \isakeyword{fixes}\ {\isasymphi}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ {\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}% +\begin{isamarkuptext}% +The second line contains the expression, which is the + merge of two partial order locales. The parameter of the second one + is \isa{le{\isacharprime}} with new infix syntax \isa{{\isasympreceq}}. The + parameters of the entire locale are \isa{le}, \isa{le{\isacharprime}} and + \isa{{\isasymphi}}. This is their \emph{canonical order}, + which is obtained by a left-to-right traversal of the expression, + where only the new parameters are appended to the end of the list. The + parameters introduced in the locale elements of the declaration + follow. + + In renamings parameters are referred to by position in the canonical + order; an underscore is used to skip a parameter position, which is + then not renamed. Renaming deletes the syntax of a parameter unless + a new mixfix annotation is given. + + Parameter renamings are morphisms between locales. These can be + lifted to terms and theorems and thus be applied to assumptions and + conclusions. The assumption of a merge is the conjunction of the + assumptions of the merged locale. The conclusions of a merge are + obtained by appending the conclusions of the left locale and of the + right locale.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +The locale \isa{order{\isacharunderscore}preserving} contains theorems for both + orders \isa{{\isasymsqsubseteq}} and \isa{{\isasympreceq}}. How can one refer to a theorem for + a particular order, \isa{{\isasymsqsubseteq}} or \isa{{\isasympreceq}}? Names in locales are + qualified by the locale parameters. More precisely, a name is + qualified by the parameters of the locale in which its declaration + occurs. Here are examples:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{context}\isamarkupfalse% +\ order{\isacharunderscore}preserving\ \isakeyword{begin}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\begin{isamarkuptext}% +\isa{le{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymsqsubseteq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}z} + + \isa{le{\isacharunderscore}le{\isacharprime}{\isacharunderscore}{\isasymphi}{\isachardot}hom{\isacharunderscore}le}: \isa{{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +When renaming a locale, the morphism is also applied + to the qualifiers. Hence theorems for the partial order \isa{{\isasympreceq}} + are qualified by \isa{le{\isacharprime}}. For example, \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}% +\ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\isanewline +\isaindent{\ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z% +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{end}\isamarkupfalse% +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\begin{isamarkuptext}% +This example reveals that there is no infix syntax for the strict + version of \isa{{\isasympreceq}}! This can, of course, not be introduced + automatically, but it can be declared manually through an abbreviation.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{abbreviation}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ order{\isacharunderscore}preserving{\isacharparenright}\isanewline +\ \ \ \ less{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymprec}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}less{\isacharprime}\ {\isasymequiv}\ partial{\isacharunderscore}order{\isachardot}less\ le{\isacharprime}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Now the theorem is displayed nicely as + \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymprec}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymprec}\ {\isacharquery}z}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Not only names of theorems are qualified. In fact, all names + are qualified, in particular names introduced by definitions and + abbreviations. The name of the strict order of \isa{{\isasymsqsubseteq}} is \isa{le{\isachardot}less} and therefore \isa{le{\isacharprime}{\isachardot}less} is the name of the strict + order of \isa{{\isasympreceq}}. Hence, the equation in the above abbreviation + could have been written as \isa{less{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}less}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Two more locales illustrate working with locale expressions. + A map \isa{{\isasymphi}} is a lattice homomorphism if it preserves meet and join.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{locale}\isamarkupfalse% +\ lattice{\isacharunderscore}hom\ {\isacharequal}\ lattice\ {\isacharplus}\ lattice\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline +\ \ \ \ \isakeyword{fixes}\ {\isasymphi}\isanewline +\ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}meet{\isacharcolon}\isanewline +\ \ \ \ \ \ \ \ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}lattice{\isachardot}meet\ le\ x\ y{\isacharparenright}\ {\isacharequal}\ lattice{\isachardot}meet\ le{\isacharprime}\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ hom{\isacharunderscore}join{\isacharcolon}\isanewline +\ \ \ \ \ \ \ \ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}lattice{\isachardot}join\ le\ x\ y{\isacharparenright}\ {\isacharequal}\ lattice{\isachardot}join\ le{\isacharprime}\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \isacommand{abbreviation}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline +\ \ \ \ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline +\ \ \isacommand{abbreviation}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline +\ \ \ \ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}join{\isachardoublequoteclose}% +\begin{isamarkuptext}% +A homomorphism is an endomorphism if both orders coincide.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{locale}\isamarkupfalse% +\ lattice{\isacharunderscore}end\ {\isacharequal}\isanewline +\ \ \ \ lattice{\isacharunderscore}hom\ le\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ le\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}% +\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. Renamings are + indicated by $\sqsubseteq \mapsto \preceq$ etc. The expression + imported by \isa{lattice{\isacharunderscore}end} identifies the first and second + parameter of \isa{lattice{\isacharunderscore}hom}. By looking at the inheritance diagram it would seem + that two identical copies of each of the locales \isa{partial{\isacharunderscore}order} and \isa{lattice} are imported. This is not the + case! Inheritance paths with identical morphisms are detected and + the conclusions of the respecitve locales appear only once. + +\begin{figure} +\hrule \vspace{2ex} +\begin{center} +\begin{tikzpicture} + \node (o) at (0,0) {\isa{partial{\isacharunderscore}order}}; + \node (oh) at (1.5,-2) {\isa{order{\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{\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{\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{interpretation}\isamarkupfalse% +\ lattice{\isacharunderscore}hom\ {\isasymsubseteq}\ order{\isacharunderscore}preserving% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ unfold{\isacharunderscore}locales\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ y\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ le{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymphi}\ y\ {\isacharequal}\ {\isacharparenleft}{\isasymphi}\ x\ {\isasymsqunion}{\isacharprime}\ {\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ hom{\isacharunderscore}join\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ le{\isacharprime}{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Theorems and other declarations --- syntax, in particular --- + from the locale \isa{order{\isacharunderscore}preserving} are now active in \isa{lattice{\isacharunderscore}hom}, for example + + \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: + \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymprec}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymprec}\ {\isacharquery}z}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\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, we + 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 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 original work of Kamm\"uller on locales \cite{KammullerEtAl1999} + may be of interest from a historical perspective. The mathematical + background on orders and lattices is taken from Jacobson's textbook + on algebra \cite[Chapter~8]{Jacobson1985}.% +\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} \\[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{rename} & ::= + & \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' \\ + \textit{expr} & ::= + & \textit{renamed-expr} ( ``\textbf{+}'' \textit{renamed-expr} )$^*$ \\ + \textit{renamed-expr} & ::= + & ( \textit{qualified-name} $|$ + ``\textbf{(}'' \textit{expr} ``\textbf{)}'' ) \textit{rename}$^*$ \\[2ex] + + \multicolumn{3}{l}{Declaration of Locales} \\ + + \textit{locale} & ::= + & \textit{element}$^+$ \\ + & | & \textit{locale-expr} [ ``\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{insts} & ::= & [ ``\textbf{[}'' \textit{term}$^+$ + ``\textbf{]}'' ] \\ + & & [ \textbf{where} \textit{equation} ( \textbf{and} + \textit{equation} )$^*$ ] \\ + \textit{toplevel} & ::= + & \textbf{interpretation} \textit{name} ( ``$<$'' $|$ + ``$\subseteq$'' ) \textit{expr} \textit{proof} \\ + & | + & \textbf{interpretation} [ \textit{attr-name} ``\textbf{:}'' ] + \textit{expr} \textit{insts} \textit{proof} \\ + & | + & \textbf{interpret} [ \textit{attr-name} ``\textbf{:}'' ] + \textit{expr} \textit{insts} \textit{proof} \\[2ex] + + \multicolumn{3}{l}{Diagnostics} \\ + + \textit{toplevel} & ::= + & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{locale} \\ + & | & \textbf{print\_locales} +\end{tabular} +\end{center} +\hrule +\caption{Syntax of Locale Commands.} +\label{tab:commands} +\end{table}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, + Christian Sternagel and Makarius Wenzel have made useful comments on + a draft of this document.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/GCD.tex diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/Locales.tex --- a/doc-src/Locales/Locales/document/Locales.tex Tue Jun 03 11:55:35 2008 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1320 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Locales}% -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -% -\isatagtheory -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isamarkupsection{Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The present text is based on~\cite{Ballarin2004a}. It was updated - for for Isabelle2005, but does not cover locale interpretation. - - Locales are an extension of the Isabelle proof assistant. They - provide support for modular reasoning. Locales were initially - developed by Kamm\"uller~\cite{Kammuller2000} to support reasoning - in abstract algebra, but are applied also in other domains --- for - example, bytecode verification~\cite{Klein2003}. - - Kamm\"uller's original design, implemented in Isabelle99, provides, in - addition to - means for declaring locales, a set of ML functions that were used - along with ML tactics in a proof. In the meantime, the input format - for proof in Isabelle has changed and users write proof - scripts in ML only rarely if at all. Two new proof styles are - available, and can - be used interchangeably: linear proof scripts that closely resemble ML - tactics, and the structured Isar proof language by - Wenzel~\cite{Wenzel2002a}. Subsequently, Wenzel re-implemented - locales for - the new proof format. The implementation, available with - Isabelle2002 or later, constitutes a complete re-design and exploits that - both Isar and locales are based on the notion of context, - and thus locales are seen as a natural extension of Isar. - Nevertheless, locales can also be used with proof scripts: - their use does not require a deep understanding of the structured - Isar proof style. - - At the same time, Wenzel considerably extended locales. The most - important addition are locale expressions, which allow to combine - locales more freely. Previously only - linear inheritance was possible. Now locales support multiple - inheritance through a normalisation algorithm. New are also - structures, which provide special syntax for locale parameters that - represent algebraic structures. - - Unfortunately, Wenzel provided only an implementation but hardly any - documentation. Besides providing documentation, the present paper - is a high-level description of locales, and in particular locale - expressions. It is meant as a first step towards the semantics of - locales, and also as a base for comparing locales with module concepts - in other provers. It also constitutes the base for future - extensions of locales in Isabelle. - The description was derived mainly by experimenting - with locales and partially also by inspecting the code. - - The main contribution of the author of the present paper is the - abstract description of Wenzel's version of locales, and in - particular of the normalisation algorithm for locale expressions (see - Section~\ref{sec-normal-forms}). Contributions to the - implementation are confined to bug fixes and to provisions that - enable the use of locales with linear proof scripts. - - Concepts are introduced along with examples, so that the text can be - used as tutorial. It is assumed that the reader is somewhat - familiar with Isabelle proof scripts. Examples have been phrased as - structured - Isar proofs. However, in order to understand the key concepts, - including locales expressions and their normalisation, detailed - knowledge of Isabelle is not necessary. - -\nocite{Nipkow2003,Wenzel2002b,Wenzel2003}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Locales: Beyond Proof Contexts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In tactic-based provers the application of a sequence of proof - tactics leads to a proof state. This state is usually hard to - predict from looking at the tactic script, unless one replays the - proof step-by-step. The structured proof language Isar is - different. It is additionally based on \emph{proof contexts}, - which are directly visible in Isar scripts, and since tactic - sequences tend to be short, this commonly leads to clearer proof - scripts. - - Goals are stated with the \textbf{theorem} - command. This is followed by a proof. When discharging a goal - requires an elaborate argument - (rather than the application of a single tactic) a new context - may be entered (\textbf{proof}). Inside the context, variables may - be fixed (\textbf{fix}), assumptions made (\textbf{assume}) and - intermediate goals stated (\textbf{have}) and proved. The - assumptions must be dischargeable by premises of the surrounding - goal, and once this goal has been proved (\textbf{show}) the proof context - can be closed (\textbf{qed}). Contexts inherit from surrounding - contexts, but it is not possible - to export from them (with exception of the proved goal); - they ``disappear'' after the closing \textbf{qed}. - Facts may have attributes --- for example, identifying them as - default to the simplifier or classical reasoner. - - Locales extend proof contexts in various ways: - \begin{itemize} - \item - Locales are usually \emph{named}. This makes them persistent. - \item - Fixed variables may have \emph{syntax}. - \item - It is possible to \emph{add} and \emph{export} facts. - \item - Locales can be combined and modified with \emph{locale - expressions}. - \end{itemize} - The Locales facility extends the Isar language: it provides new ways - of stating and managing facts, but it does not modify the language - for proofs. Its purpose is to support writing modular proofs.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Simple Locales% -} -\isamarkuptrue% -% -\isamarkupsubsection{Syntax and Terminology% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The grammar of Isar is extended by commands for locales as shown in - Figure~\ref{fig-grammar}. - A key concept, introduced by Wenzel, is that - locales are (internally) lists - of \emph{context elements}. There are five kinds, identified - by the keywords \textbf{fixes}, \textbf{constrains}, - \textbf{assumes}, \textbf{defines} and \textbf{notes}. - - \begin{figure} - \hrule - \vspace{2ex} - \begin{small} - \begin{tabular}{l>$c<$l} - \textit{attr-name} & ::= - & \textit{name} $|$ \textit{attribute} $|$ - \textit{name} \textit{attribute} \\ - - \textit{locale-expr} & ::= - & \textit{locale-expr1} ( ``\textbf{+}'' \textit{locale-expr1} )$^*$ \\ - \textit{locale-expr1} & ::= - & ( \textit{qualified-name} $|$ - ``\textbf{(}'' \textit{locale-expr} ``\textbf{)}'' ) \\ - & & ( \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' )$^*$ \\ - - \textit{fixes} & ::= - & \textit{name} [ ``\textbf{::}'' \textit{type} ] - [ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ - \textit{mixfix} ] \\ - \textit{constrains} & ::= - & \textit{name} ``\textbf{::}'' \textit{type} \\ - \textit{assumes} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ - \textit{defines} & ::= - & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ - \textit{notes} & ::= - & [ \textit{attr-name} ``\textbf{=}'' ] - ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ - - \textit{element} & ::= - & \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ - & | - & \textbf{constrains} \textit{constrains} - ( \textbf{and} \textit{constrains} )$^*$ \\ - & | - & \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\ - & | - & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ - & | - & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ - \textit{element1} & ::= - & \textit{element} \\ - & | & \textbf{includes} \textit{locale-expr} \\ - - \textit{locale} & ::= - & \textit{element}$^+$ \\ - & | & \textit{locale-expr} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ - - \textit{in-target} & ::= - & ``\textbf{(}'' \textbf{in} \textit{qualified-name} ``\textbf{)}'' \\ - - \textit{theorem} & ::= & ( \textbf{theorem} $|$ \textbf{lemma} $|$ - \textbf{corollary} ) [ \textit{in-target} ] [ \textit{attr-name} ] \\ - - \textit{theory-level} & ::= & \ldots \\ - & | & \textbf{locale} \textit{name} [ ``\textbf{=}'' - \textit{locale} ] \\ - % note: legacy "locale (open)" omitted. - & | & ( \textbf{theorems} $|$ \textbf{lemmas} ) \\ - & & [ \textit{in-target} ] [ \textit{attr-name} ``\textbf{=}'' ] - ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ - & | & \textbf{declare} [ \textit{in-target} ] ( \textit{qualified-name} - [ \textit{attribute} ] )$^+$ \\ - & | & \textit{theorem} \textit{proposition} \textit{proof} \\ - & | & \textit{theorem} \textit{element1}$^*$ - \textbf{shows} \textit{proposition} \textit{proof} \\ - & | & \textbf{print\_locale} \textit{locale} \\ - & | & \textbf{print\_locales} - \end{tabular} - \end{small} - \vspace{2ex} - \hrule - \caption{Locales extend the grammar of Isar.} - \label{fig-grammar} - \end{figure} - - At the theory level --- that is, at the outer syntactic level of an - Isabelle input file --- \textbf{locale} declares a named - locale. Other kinds of locales, - locale expressions and unnamed locales, will be introduced later. When - declaring a named locale, it is possible to \emph{import} another - named locale, or indeed several ones by importing a locale - expression. The second part of the declaration, also optional, - consists of a number of context element declarations. - - A number of Isar commands have an additional, optional \emph{target} - argument, which always refers to a named locale. These commands - are \textbf{theorem} (together with \textbf{lemma} and - \textbf{corollary}), \textbf{theorems} (and - \textbf{lemmas}), and \textbf{declare}. The effect of specifying a target is - that these commands focus on the specified locale, not the - surrounding theory. Commands that are used to - prove new theorems will add them not to the theory, but to the - locale. Similarly, \textbf{declare} modifies attributes of theorems - that belong to the specified target. Additionally, for - \textbf{theorem} (and related commands), theorems stored in the target - can be used in the associated proof scripts. - - The Locales package provides a \emph{long goals format} for - propositions stated with \textbf{theorem} (and friends). While - normally a goal is just a formula, a long goal is a list of context - elements, followed by the keyword \textbf{shows}, followed by the - formula. Roughly speaking, the context elements are - (additional) premises. For an example, see - Section~\ref{sec-includes}. The list of context elements in a long goal - is also called \emph{unnamed locale}. - - Finally, there are two commands to inspect locales when working in - interactive mode: \textbf{print\_locales} prints the names of all - targets - visible in the current theory, \textbf{print\_locale} outputs the - elements of a named locale or locale expression. - - The following presentation will use notation of - Isabelle's meta logic, hence a few sentences to explain this. - The logical - primitives are universal quantification (\isa{{\isasymAnd}}), entailment - (\isa{{\isasymLongrightarrow}}) and equality (\isa{{\isasymequiv}}). Variables (not bound - variables) are sometimes preceded by a question mark. The logic is - typed. Type variables are denoted by \isa{{\isacharprime}a}, \isa{{\isacharprime}b} - etc., and \isa{{\isasymRightarrow}} is the function type. Double brackets \isa{{\isasymlbrakk}} and \isa{{\isasymrbrakk}} are used to abbreviate nested entailment.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Parameters, Assumptions and Facts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -From a logical point of view a \emph{context} is a formula schema of - the form -\[ - \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}{\isasymdots}x\isactrlsub n{\isachardot}\ {\isasymlbrakk}\ C\isactrlsub {\isadigit{1}}{\isacharsemicolon}\ {\isasymdots}\ {\isacharsemicolon}C\isactrlsub m\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymdots}} -\] - The variables $\isa{x\isactrlsub {\isadigit{1}}}, \ldots, \isa{x\isactrlsub n}$ are - called \emph{parameters}, the premises $\isa{C\isactrlsub {\isadigit{1}}}, \ldots, - \isa{C\isactrlsub n}$ \emph{assumptions}. A formula \isa{F} - holds in this context if -\begin{equation} -\label{eq-fact-in-context} - \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}{\isasymdots}x\isactrlsub n{\isachardot}\ {\isasymlbrakk}\ C\isactrlsub {\isadigit{1}}{\isacharsemicolon}\ {\isasymdots}\ {\isacharsemicolon}C\isactrlsub m\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ F} -\end{equation} - is valid. The formula is called a \emph{fact} of the context. - - A locale allows fixing the parameters \isa{x\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ x\isactrlsub n} and making the assumptions \isa{C\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ C\isactrlsub m}. This implicitly builds the context in - which the formula \isa{F} can be established. - Parameters of a locale correspond to the context element - \textbf{fixes}, and assumptions may be declared with - \textbf{assumes}. Using these context elements one can define - the specification of semigroups.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ semi\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ prod\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymcdot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymcdot}\ y{\isacharparenright}\ {\isasymcdot}\ z\ {\isacharequal}\ x\ {\isasymcdot}\ {\isacharparenleft}y\ {\isasymcdot}\ z{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The parameter \isa{prod} has a - syntax annotation enabling the infix ``\isa{{\isasymcdot}}'' in the - assumption of associativity. Parameters may have arbitrary mixfix - syntax, like constants. In the example, the type of \isa{prod} is - specified explicitly. This is not necessary. If no type is - specified, a most general type is inferred simultaneously for all - parameters, taking into account all assumptions (and type - specifications of parameters, if present).% -\footnote{Type inference also takes into account type constraints, - definitions and import, as introduced later.} - - Free variables in assumptions are implicitly universally quantified, - unless they are parameters. Hence the context defined by the locale - \isa{semi} is -\[ - \isa{{\isasymAnd}prod{\isachardot}\ {\isasymlbrakk}\ {\isasymAnd}x\ y\ z{\isachardot}\ prod\ {\isacharparenleft}prod\ x\ y{\isacharparenright}\ z\ {\isacharequal}\ prod\ x\ {\isacharparenleft}prod\ y\ z{\isacharparenright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymdots}} -\] - The locale can be extended to commutative semigroups.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ comm{\isacharunderscore}semi\ {\isacharequal}\ semi\ {\isacharplus}\isanewline -\ \ \isakeyword{assumes}\ comm{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymcdot}\ y\ {\isacharequal}\ y\ {\isasymcdot}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -This locale \emph{imports} all elements of \isa{semi}. The - latter locale is called the import of \isa{comm{\isacharunderscore}semi}. The - definition adds commutativity, hence its context is -\begin{align*% -} - \isa{{\isasymAnd}prod{\isachardot}\ {\isasymlbrakk}} & - \isa{{\isasymAnd}x\ y\ z{\isachardot}\ prod\ {\isacharparenleft}prod\ x\ y{\isacharparenright}\ z\ {\isacharequal}\ prod\ x\ {\isacharparenleft}prod\ y\ z{\isacharparenright}{\isacharsemicolon}} \\ - & \isa{{\isasymAnd}x\ y{\isachardot}\ prod\ x\ y\ {\isacharequal}\ prod\ y\ x\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymdots}} -\end{align*% -} - One may now derive facts --- for example, left-commutativity --- in - the context of \isa{comm{\isacharunderscore}semi} by specifying this locale as - target, and by referring to the names of the assumptions \isa{assoc} and \isa{comm} in the proof.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ comm{\isacharunderscore}semi{\isacharparenright}\ lcomm{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isasymcdot}\ {\isacharparenleft}y\ {\isasymcdot}\ z{\isacharparenright}\ {\isacharequal}\ y\ {\isasymcdot}\ {\isacharparenleft}x\ {\isasymcdot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymcdot}\ {\isacharparenleft}y\ {\isasymcdot}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymcdot}\ y{\isacharparenright}\ {\isasymcdot}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ {\isacharparenleft}y\ {\isasymcdot}\ x{\isacharparenright}\ {\isasymcdot}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ comm{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ y\ {\isasymcdot}\ {\isacharparenleft}x\ {\isasymcdot}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -In this equational Isar proof, ``\isa{{\isasymdots}}'' refers to the - right hand side of the preceding equation. - After the proof is finished, the fact \isa{lcomm} is added to - the locale \isa{comm{\isacharunderscore}semi}. This is done by adding a - \textbf{notes} element to the internal representation of the locale, - as explained the next section.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Locale Predicates and the Internal Representation of - Locales% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{sec-locale-predicates} - In mathematical texts, often arbitrary but fixed objects with - certain properties are considered --- for instance, an arbitrary but - fixed group $G$ --- with the purpose of establishing facts valid for - any group. These facts are subsequently used on other objects that - also have these properties. - - Locales permit the same style of reasoning. Exporting a fact $F$ - generalises the fixed parameters and leads to a (valid) formula of the - form of equation~(\ref{eq-fact-in-context}). If a locale has many - assumptions - (possibly accumulated through a number of imports) this formula can - become large and cumbersome. Therefore, Wenzel introduced - predicates that abbreviate the assumptions of locales. These - predicates are not confined to the locale but are visible in the - surrounding theory. - - The definition of the locale \isa{semi} generates the \emph{locale - predicate} \isa{semi} over the type of the parameter \isa{prod}, - hence the predicate's type is \isa{{\isacharparenleft}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ bool}. Its - definition is -\begin{equation} - \tag*{\isa{semi{\isacharunderscore}def}:} \isa{semi\ {\isacharquery}prod\ {\isasymequiv}\ {\isasymforall}x\ y\ z{\isachardot}\ {\isacharquery}prod\ {\isacharparenleft}{\isacharquery}prod\ x\ y{\isacharparenright}\ z\ {\isacharequal}\ {\isacharquery}prod\ x\ {\isacharparenleft}{\isacharquery}prod\ y\ z{\isacharparenright}}. -\end{equation} - In the case where the locale has no import, the generated - predicate abbreviates all assumptions and is over the parameters - that occur in these assumptions. - - The situation is more complicated when a locale extends - another locale, as is the case for \isa{comm{\isacharunderscore}semi}. Two - predicates are defined. The predicate - \isa{comm{\isacharunderscore}semi{\isacharunderscore}axioms} corresponds to the new assumptions and is - called \emph{delta predicate}, the locale - predicate \isa{comm{\isacharunderscore}semi} captures the content of all the locale, - including the import. - If a locale has neither assumptions nor import, no predicate is - defined. If a locale has import but no assumptions, only the locale - predicate is defined.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\begin{isamarkuptext}% -The Locales package generates a number of theorems for locale and - delta predicates. All predicates have a definition and an - introduction rule. Locale predicates that are defined in terms of - other predicates (which is the case if and only if the locale has - import) also have a number of elimination rules (called - \emph{axioms}). All generated theorems for the predicates of the - locales \isa{semi} and \isa{comm{\isacharunderscore}semi} are shown in - Figures~\ref{fig-theorems-semi} and~\ref{fig-theorems-comm-semi}, - respectively. - \begin{figure} - \hrule - \vspace{2ex} - Theorems generated for the predicate \isa{semi}. - \begin{gather} - \tag*{\isa{semi{\isacharunderscore}def}:} \isa{semi\ {\isacharquery}prod\ {\isasymequiv}\ {\isasymforall}x\ y\ z{\isachardot}\ {\isacharquery}prod\ {\isacharparenleft}{\isacharquery}prod\ x\ y{\isacharparenright}\ z\ {\isacharequal}\ {\isacharquery}prod\ x\ {\isacharparenleft}{\isacharquery}prod\ y\ z{\isacharparenright}} \\ - \tag*{\isa{semi{\isachardot}intro}:} \isa{{\isacharparenleft}{\isasymAnd}x\ y\ z{\isachardot}\ {\isacharquery}prod\ {\isacharparenleft}{\isacharquery}prod\ x\ y{\isacharparenright}\ z\ {\isacharequal}\ {\isacharquery}prod\ x\ {\isacharparenleft}{\isacharquery}prod\ y\ z{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ semi\ {\isacharquery}prod} - \end{gather} - \hrule - \caption{Theorems for the locale predicate \isa{semi}.} - \label{fig-theorems-semi} - \end{figure} - - \begin{figure}[h] - \hrule - \vspace{2ex} - Theorems generated for the predicate \isa{comm{\isacharunderscore}semi{\isacharunderscore}axioms}. - \begin{gather} - \tag*{\isa{comm{\isacharunderscore}semi{\isacharunderscore}axioms{\isacharunderscore}def}:} \isa{comm{\isacharunderscore}semi{\isacharunderscore}axioms\ {\isacharquery}prod\ {\isasymequiv}\ {\isasymforall}x\ y{\isachardot}\ {\isacharquery}prod\ x\ y\ {\isacharequal}\ {\isacharquery}prod\ y\ x} \\ - \tag*{\isa{comm{\isacharunderscore}semi{\isacharunderscore}axioms{\isachardot}intro}:} \isa{{\isacharparenleft}{\isasymAnd}x\ y{\isachardot}\ {\isacharquery}prod\ x\ y\ {\isacharequal}\ {\isacharquery}prod\ y\ x{\isacharparenright}\ {\isasymLongrightarrow}\ comm{\isacharunderscore}semi{\isacharunderscore}axioms\ {\isacharquery}prod} - \end{gather} - Theorems generated for the predicate \isa{comm{\isacharunderscore}semi}. - \begin{gather} - \tag*{\isa{comm{\isacharunderscore}semi{\isacharunderscore}def}:} \isa{comm{\isacharunderscore}semi\ {\isacharquery}prod\ {\isasymequiv}\ semi\ {\isacharquery}prod\ {\isasymand}\ comm{\isacharunderscore}semi{\isacharunderscore}axioms\ {\isacharquery}prod} \\ - \tag*{\isa{comm{\isacharunderscore}semi{\isachardot}intro}:} \isa{{\isasymlbrakk}semi\ {\isacharquery}prod{\isacharsemicolon}\ comm{\isacharunderscore}semi{\isacharunderscore}axioms\ {\isacharquery}prod{\isasymrbrakk}\ {\isasymLongrightarrow}\ comm{\isacharunderscore}semi\ {\isacharquery}prod} \\ - \tag*{\isa{comm{\isacharunderscore}semi{\isachardot}axioms}:} \mbox{} \\ - \notag \isa{comm{\isacharunderscore}semi\ {\isacharquery}prod\ {\isasymLongrightarrow}\ semi\ {\isacharquery}prod} \\ - \notag \isa{comm{\isacharunderscore}semi\ {\isacharquery}prod\ {\isasymLongrightarrow}\ comm{\isacharunderscore}semi{\isacharunderscore}axioms\ {\isacharquery}prod} - \end{gather} - \hrule - \caption{Theorems for the predicates \isa{comm{\isacharunderscore}semi{\isacharunderscore}axioms} and - \isa{comm{\isacharunderscore}semi}.} - \label{fig-theorems-comm-semi} - \end{figure} - Note that the theorems generated by a locale - definition may be inspected immediately after the definition in the - Proof General interface \cite{Aspinall2000} of Isabelle through - the menu item ``Isabelle/Isar$>$Show me $\ldots>$Theorems''. - - Locale and delta predicates are used also in the internal - representation of locales as lists of context elements. While all - \textbf{fixes} in a - declaration generate internal \textbf{fixes}, all assumptions - of one locale declaration contribute to one internal \textbf{assumes} - element. The internal representation of \isa{semi} is -\[ -\begin{array}{ll} - \textbf{fixes} & \isa{prod} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}} - (\textbf{infixl} \isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}} 70) \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{assoc}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z\ {\isacharequal}\ {\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} -\end{array} -\] - and the internal representation of \isa{{\isachardoublequote}comm{\isacharunderscore}semi{\isachardoublequote}} is -\begin{equation} -\label{eq-comm-semi} -\begin{array}{ll} - \textbf{fixes} & \isa{prod} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}} - ~(\textbf{infixl}~\isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}}~70) \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{assoc}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z\ {\isacharequal}\ {\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} \\ - \textbf{assumes} & \isa{{\isachardoublequote}comm{\isacharunderscore}semi{\isacharunderscore}axioms\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{comm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}x{\isachardoublequote}} \\ - \textbf{notes} & \isa{lcomm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} -\end{array} -\end{equation} - The \textbf{notes} elements store facts of the - locales. The facts \isa{assoc} and \isa{comm} were added - during the declaration of the locales. They stem from assumptions, - which are trivially facts. The fact \isa{lcomm} was - added later, after finishing the proof in the respective - \textbf{theorem} command above. - - By using \textbf{notes} in a declaration, facts can be added - to a locale directly. Of course, these must be theorems. - Typical use of this feature includes adding theorems that are not - usually used as a default rewrite rules by the simplifier to the - simpset of the locale by a \textbf{notes} element with the attribute - \isa{{\isacharbrackleft}simp{\isacharbrackright}}. This way it is also possible to add specialised - versions of - theorems to a locale by instantiating locale parameters for unknowns - or locale assumptions for premises.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Definitions were available in Kamm\"uller's version of Locales, and - they are in Wenzel's. - The context element \textbf{defines} adds a definition of the form - \isa{p\ x\isactrlsub {\isadigit{1}}\ {\isasymdots}\ x\isactrlsub n\ {\isasymequiv}\ t} as an assumption, where \isa{p} is a - parameter of the locale (possibly an imported parameter), and \isa{t} a term that may contain the \isa{x\isactrlsub i}. The parameter may - neither occur in a previous \textbf{assumes} or \textbf{defines} - element, nor on the right hand side of the definition. Hence - recursion is not allowed. - The parameter may, however, occur in subsequent \textbf{assumes} and - on the right hand side of subsequent \textbf{defines}. We call - \isa{p} \emph{defined parameter}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ semi{\isadigit{2}}\ {\isacharequal}\ semi\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ rprod\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymodot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{defines}\ rprod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}rprod\ x\ y\ {\isasymequiv}\ y\ {\isasymcdot}\ x\ {\isachardoublequoteclose}% -\begin{isamarkuptext}% -This locale extends \isa{semi} by a second binary operation \isa{{\isachardoublequote}{\isasymodot}{\isachardoublequote}} that is like \isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}} but with - reversed arguments. The - definition of the locale generates the predicate \isa{semi{\isadigit{2}}}, - which is equivalent to \isa{semi}, but no \isa{semi{\isadigit{2}}{\isacharunderscore}axioms}. - The difference between \textbf{assumes} and \textbf{defines} lies in - the way parameters are treated on export.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Export% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A fact is exported out - of a locale by generalising over the parameters and adding - assumptions as premises. For brevity of the exported theorems, - locale predicates are used. Exported facts are referenced by - writing qualified names consisting of the locale name and the fact name --- - for example, -\begin{equation} - \tag*{\isa{semi{\isachardot}assoc}:} \isa{semi\ {\isacharquery}prod\ {\isasymLongrightarrow}\ {\isacharquery}prod\ {\isacharparenleft}{\isacharquery}prod\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharquery}z\ {\isacharequal}\ {\isacharquery}prod\ {\isacharquery}x\ {\isacharparenleft}{\isacharquery}prod\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}}. -\end{equation} - Defined parameters receive special treatment. Instead of adding a - premise for the definition, the definition is unfolded in the - exported theorem. In order to illustrate this we prove that the - reverse operation \isa{{\isachardoublequote}{\isasymodot}{\isachardoublequote}} defined in the locale - \isa{semi{\isadigit{2}}} is also associative.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ semi{\isadigit{2}}{\isacharparenright}\ r{\isacharunderscore}assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ rprod{\isacharunderscore}def\ assoc{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -The exported fact is -\begin{equation} - \tag*{\isa{semi{\isadigit{2}}{\isachardot}r{\isacharunderscore}assoc}:} \isa{semi{\isadigit{2}}\ {\isacharquery}prod\ {\isasymLongrightarrow}\ {\isacharquery}prod\ {\isacharquery}z\ {\isacharparenleft}{\isacharquery}prod\ {\isacharquery}y\ {\isacharquery}x{\isacharparenright}\ {\isacharequal}\ {\isacharquery}prod\ {\isacharparenleft}{\isacharquery}prod\ {\isacharquery}z\ {\isacharquery}y{\isacharparenright}\ {\isacharquery}x}. -\end{equation} - The defined parameter is not present but is replaced by its - definition. Note that the definition itself is not exported, hence - there is no \isa{semi{\isadigit{2}}{\isachardot}rprod{\isacharunderscore}def}.% -\footnote{The definition could alternatively be exported using a - let-construct if there was one in Isabelle's meta-logic. Let is - usually defined in object-logics.}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Locale Expressions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locale expressions provide a simple language for combining - locales. They are an effective means of building complex - specifications from simple ones. Locale expressions are the main - innovation of the version of Locales discussed here. Locale - expressions are also reason for introducing locale predicates.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Rename and Merge% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The grammar of locale expressions is part of the grammar in - Figure~\ref{fig-grammar}. Locale names are locale - expressions, and - further expressions are obtained by \emph{rename} and \emph{merge}. -\begin{description} -\item[Rename.] - The locale expression $e\: q_1 \ldots q_n$ denotes - the locale of $e$ where pa\-ra\-me\-ters, in the order in - which they are fixed, are renamed to $q_1$ to $q_n$. - The expression is only well-formed if $n$ does not - exceed the number of parameters of $e$. Underscores denote - parameters that are not renamed. - Renaming by default removes mixfix syntax, but new syntax may be - specified. -\item[Merge.] - The locale expression $e_1 + e_2$ denotes - the locale obtained by merging the locales of $e_1$ - and $e_2$. This locale contains the context elements - of $e_1$, followed by the context elements of $e_2$. - - In actual fact, the semantics of the merge operation - is more complicated. If $e_1$ and $e_2$ are expressions containing - the same name, followed by - identical parameter lists, then the merge of both will contain - the elements of those locales only once. Details are explained in - Section~\ref{sec-normal-forms} below. - - The merge operation is associative but not commutative. The latter - is because parameters of $e_1$ appear - before parameters of $e_2$ in the composite expression. -\end{description} - - Rename can be used if a different parameter name seems more - appropriate --- for example, when moving from groups to rings, a - parameter \isa{G} representing the group could be changed to - \isa{R}. Besides of this stylistic use, renaming is important in - combination with merge. Both operations are used in the following - specification of semigroup homomorphisms.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ semi{\isacharunderscore}hom\ {\isacharequal}\ comm{\isacharunderscore}semi\ sum\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymoplus}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{5}}{\isacharparenright}\ {\isacharplus}\ comm{\isacharunderscore}semi\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ hom\isanewline -\ \ \isakeyword{assumes}\ hom{\isacharcolon}\ {\isachardoublequoteopen}hom\ {\isacharparenleft}x\ {\isasymoplus}\ y{\isacharparenright}\ {\isacharequal}\ hom\ x\ {\isasymcdot}\ hom\ y{\isachardoublequoteclose}% -\begin{isamarkuptext}% -This locale defines a context with three parameters \isa{sum}, - \isa{prod} and \isa{hom}. The first two parameters have - mixfix syntax. They are associative operations, - the first of type \isa{{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a}, the second of - type \isa{{\isacharbrackleft}{\isacharprime}b{\isacharcomma}\ {\isacharprime}b{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}b}. - - How are facts that are imported via a locale expression identified? - Facts are always introduced in a named locale (either in the - locale's declaration, or by using the locale as target in - \textbf{theorem}), and their names are - qualified by the parameter names of this locale. - Hence the full name of associativity in \isa{semi} is - \isa{prod{\isachardot}assoc}. Renaming parameters of a target also renames - the qualifier of facts. Hence, associativity of \isa{sum} is - \isa{sum{\isachardot}assoc}. Several parameters are separated by - underscores in qualifiers. For example, the full name of the fact - \isa{hom} in the locale \isa{semi{\isacharunderscore}hom} is \isa{sum{\isacharunderscore}prod{\isacharunderscore}hom{\isachardot}hom}. - - The following example is quite artificial, it illustrates the use of - facts, though.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ semi{\isacharunderscore}hom{\isacharparenright}\ {\isachardoublequoteopen}hom\ x\ {\isasymcdot}\ {\isacharparenleft}hom\ y\ {\isasymcdot}\ hom\ z{\isacharparenright}\ {\isacharequal}\ hom\ {\isacharparenleft}x\ {\isasymoplus}\ {\isacharparenleft}y\ {\isasymoplus}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}hom\ x\ {\isasymcdot}\ {\isacharparenleft}hom\ y\ {\isasymcdot}\ hom\ z{\isacharparenright}\ {\isacharequal}\ hom\ y\ {\isasymcdot}\ {\isacharparenleft}hom\ x\ {\isasymcdot}\ hom\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ prod{\isachardot}lcomm{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ hom\ {\isacharparenleft}y\ {\isasymoplus}\ {\isacharparenleft}x\ {\isasymoplus}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ hom{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ hom\ {\isacharparenleft}x\ {\isasymoplus}\ {\isacharparenleft}y\ {\isasymoplus}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ sum{\isachardot}lcomm{\isacharparenright}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Importing via a locale expression imports all facts of - the imported locales, hence both \isa{sum{\isachardot}lcomm} and \isa{prod{\isachardot}lcomm} are - available in \isa{hom{\isacharunderscore}semi}. The import is dynamic --- that is, - whenever facts are added to a locale, they automatically - become available in subsequent \textbf{theorem} commands that use - the locale as a target, or a locale importing the locale.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Normal Forms% -} -\isamarkuptrue% -% -\label{sec-normal-forms} -\newcommand{\I}{\mathcal{I}} -\newcommand{\F}{\mathcal{F}} -\newcommand{\N}{\mathcal{N}} -\newcommand{\C}{\mathcal{C}} -\newcommand{\App}{\mathbin{\overline{@}}} -% -\begin{isamarkuptext}% -Locale expressions are interpreted in a two-step process. First, an - expression is normalised, then it is converted to a list of context - elements. - - Normal forms are based on \textbf{locale} declarations. These - consist of an import section followed by a list of context - elements. Let $\I(l)$ denote the locale expression imported by - locale $l$. If $l$ has no import then $\I(l) = \varepsilon$. - Likewise, let $\F(l)$ denote the list of context elements, also - called the \emph{context fragment} of $l$. Note that $\F(l)$ - contains only those context elements that are stated in the - declaration of $l$, not imported ones. - -\paragraph{Example 1.} Consider the locales \isa{semi} and \isa{comm{\isacharunderscore}semi}. We have $\I(\isa{semi}) = \varepsilon$ and - $\I(\isa{comm{\isacharunderscore}semi}) = \isa{semi}$, and the context fragments - are -\begin{align*% -} - \F(\isa{semi}) & = \left[ -\begin{array}{ll} - \textbf{fixes} & \isa{prod} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}} - ~(\textbf{infixl}~\isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}}~70) \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{assoc}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z\ {\isacharequal}\ {\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} -\end{array} \right], \\ - \F(\isa{comm{\isacharunderscore}semi}) & = \left[ -\begin{array}{ll} - \textbf{assumes} & \isa{{\isachardoublequote}comm{\isacharunderscore}semi{\isacharunderscore}axioms\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{comm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}x{\isachardoublequote}} \\ - \textbf{notes} & \isa{lcomm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} -\end{array} \right]. -\end{align*% -} - Let $\pi_0(\F(l))$ denote the list of parameters defined in the - \textbf{fixes} elements of $\F(l)$ in the order of their occurrence. - The list of parameters of a locale expression $\pi(e)$ is defined as - follows: -\begin{align*% -} - \pi(l) & = \pi(\I(l)) \App \pi_0(\F(l)) \text{, for named locale $l$.} \\ - \pi(e\: q_1 \ldots q_n) & = \text{$[q_1, \ldots, q_n, p_{n+1}, \ldots, - p_{m}]$, where $\pi(e) = [p_1, \ldots, p_m]$.} \\ - \pi(e_1 + e_2) & = \pi(e_1) \App \pi(e_2) -\end{align*% -} - The operation $\App$ concatenates two lists but omits elements from - the second list that are also present in the first list. - It is not possible to rename more parameters than there - are present in an expression --- that is, $n \le m$ --- otherwise - the renaming is illegal. If $q_i - = \_$ then the $i$th entry of the resulting list is $p_i$. - - In the normalisation phase, imports of named locales are unfolded, and - renames and merges are recursively propagated to the imported locale - expressions. The result is a list of locale names, each with a full - list of parameters, where locale names occurring with the same parameter - list twice are removed. Let $\N$ denote normalisation. It is defined - by these equations: -\begin{align*% -} - \N(l) & = \N(\I(l)) \App [l\:\pi(l)] \text{, for named locale $l$.} \\ - \N(e\: q_1 \ldots q_n) & = \N(e)\:[q_1 \ldots q_n/\pi(e)] \\ - \N(e_1 + e_2) & = \N(e_1) \App \N(e_2) -\end{align*% -} - Normalisation yields a list of \emph{identifiers}. An identifier - consists of a locale name and a (possibly empty) list of parameters. - - In the second phase, the list of identifiers $\N(e)$ is converted to - a list of context elements $\C(e)$ by converting each identifier to - a list of context elements, and flattening the obtained list. - Conversion of the identifier $l\:q_1 \ldots q_n$ yields the list of - context elements $\F(l)$, but with the following modifications: -\begin{itemize} -\item - Rename the parameter in the $i$th \textbf{fixes} element of $\F(l)$ - to $q_i$, $i = 1, \ldots, n$. If the parameter name is actually - changed then delete the syntax annotation. Renaming a parameter may - also change its type. -\item - Perform the same renamings on all occurrences of parameters (fixed - variables) in \textbf{assumes}, \textbf{defines} and \textbf{notes} - elements. -\item - Qualify names of facts by $q_1\_\ldots\_q_n$. -\end{itemize} - The locale expression is \emph{well-formed} if it contains no - illegal renamings and the following conditions on $\C(e)$ hold, - otherwise the expression is rejected: -\begin{itemize} -\item Parameters in \textbf{fixes} are distinct; -\item Free variables in \textbf{assumes} and - \textbf{defines} occur in preceding \textbf{fixes};% -\footnote{This restriction is relaxed for contexts obtained with - \textbf{includes}, see Section~\ref{sec-includes}.} -\item Parameters defined in \textbf{defines} must neither occur in - preceding \textbf{assumes} nor \textbf{defines}. -\end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Examples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\paragraph{Example 2.} - We obtain the context fragment $\C(\isa{comm{\isacharunderscore}semi})$ of the - locale \isa{comm{\isacharunderscore}semi}. First, the parameters are computed. -\begin{align*% -} - \pi(\isa{semi}) & = [\isa{prod}] \\ - \pi(\isa{comm{\isacharunderscore}semi}) & = \pi(\isa{semi}) \App [] - = [\isa{prod}] -\end{align*% -} - Next, the normal form of the locale expression - \isa{comm{\isacharunderscore}semi} is obtained. -\begin{align*% -} - \N(\isa{semi}) & = [\isa{semi} \isa{prod}] \\ - \N(\isa{comm{\isacharunderscore}semi}) & = \N(\isa{semi}) \App - [\isa{comm{\isacharunderscore}semi\ prod}] - = [\isa{semi\ prod}, \isa{comm{\isacharunderscore}semi\ prod}] -\end{align*% -} - Converting this to a list of context elements leads to the - list~(\ref{eq-comm-semi}) shown in - Section~\ref{sec-locale-predicates}, but with fact names qualified - by \isa{prod} --- for example, \isa{prod{\isachardot}assoc}. - Qualification was omitted to keep the presentation simple. - Isabelle's scoping rules identify the most recent fact with - qualified name $x.a$ when a fact with name $a$ is requested. - -\paragraph{Example 3.} - The locale expression \isa{comm{\isacharunderscore}semi\ sum} involves - renaming. Computing parameters yields $\pi(\isa{comm{\isacharunderscore}semi\ sum}) - = [\isa{sum}]$, the normal form is -\begin{align*% -} - \N(\isa{comm{\isacharunderscore}semi\ sum}) & = - \N(\isa{comm{\isacharunderscore}semi})[\isa{sum}/\isa{prod}] = - [\isa{semi\ sum}, \isa{comm{\isacharunderscore}semi\ sum}] -\end{align*% -} - and the list of context elements -\[ -\begin{array}{ll} - \textbf{fixes} & \isa{sum} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}} - ~(\textbf{infixl}~\isa{{\isachardoublequote}{\isasymoplus}{\isachardoublequote}}~65) \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi\ sum{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isachardot}assoc}: \isa{{\isachardoublequote}{\isacharparenleft}{\isacharquery}x\ {\isasymoplus}\ {\isacharquery}y{\isacharparenright}\ {\isasymoplus}\ {\isacharquery}z\ {\isacharequal}\ sum\ {\isacharquery}x\ {\isacharparenleft}sum\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} \\ - \textbf{assumes} & \isa{{\isachardoublequote}comm{\isacharunderscore}semi{\isacharunderscore}axioms\ sum{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isachardot}comm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymoplus}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}y\ {\isasymoplus}\ {\isacharquery}x{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isachardot}lcomm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymoplus}\ {\isacharparenleft}{\isacharquery}y\ {\isasymoplus}\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ {\isacharquery}y\ {\isasymoplus}\ {\isacharparenleft}{\isacharquery}x\ {\isasymoplus}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} -\end{array} -\] - -\paragraph{Example 4.} - The context defined by the locale \isa{semi{\isacharunderscore}hom} involves - merging two copies of \isa{comm{\isacharunderscore}semi}. We obtain the parameter - list and normal form: -\begin{align*% -} - \pi(\isa{semi{\isacharunderscore}hom}) & = \pi(\isa{comm{\isacharunderscore}semi\ sum} + - \isa{comm{\isacharunderscore}semi}) \App [\isa{hom}] \\ - & = (\pi(\isa{comm{\isacharunderscore}semi\ sum}) \App \pi(\isa{comm{\isacharunderscore}semi})) - \App [\isa{hom}] \\ - & = ([\isa{sum}] \App [\isa{prod}]) \App [\isa{hom}] - = [\isa{sum}, \isa{prod}, \isa{hom}] \\ - \N(\isa{semi{\isacharunderscore}hom}) & = - \N(\isa{comm{\isacharunderscore}semi\ sum} + \isa{comm{\isacharunderscore}semi}) \App \\ - & \phantom{==} - [\isa{semi{\isacharunderscore}hom\ sum\ prod\ hom}] \\ - & = (\N(\isa{comm{\isacharunderscore}semi\ sum}) \App \N(\isa{comm{\isacharunderscore}semi})) \App \\ - & \phantom{==} - [\isa{semi{\isacharunderscore}hom\ sum\ prod\ hom}] \\ - & = ([\isa{semi\ sum}, \isa{comm{\isacharunderscore}semi\ sum}] \App %\\ -% & \phantom{==} - [\isa{semi\ prod}, \isa{comm{\isacharunderscore}semi\ prod}]) \App \\ - & \phantom{==} - [\isa{semi{\isacharunderscore}hom\ sum\ prod\ hom}] \\ - & = [\isa{semi\ sum}, \isa{comm{\isacharunderscore}semi\ sum}, - \isa{semi\ prod}, \isa{comm{\isacharunderscore}semi\ prod}, \\ - & \phantom{==} - \isa{semi{\isacharunderscore}hom\ sum\ prod\ hom}]. -\end{align*% -} - Hence $\C(\isa{semi{\isacharunderscore}hom})$, shown below, is again well-formed. -\[ -\begin{array}{ll} - \textbf{fixes} & \isa{sum} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}} - ~(\textbf{infixl}~\isa{{\isachardoublequote}{\isasymoplus}{\isachardoublequote}}~65) \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi\ sum{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isachardot}assoc}: \isa{{\isachardoublequote}{\isacharparenleft}{\isacharquery}x\ {\isasymoplus}\ {\isacharquery}y{\isacharparenright}\ {\isasymoplus}\ {\isacharquery}z\ {\isacharequal}\ {\isacharquery}x\ {\isasymoplus}\ {\isacharparenleft}{\isacharquery}y\ {\isasymoplus}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} \\ - \textbf{assumes} & \isa{{\isachardoublequote}comm{\isacharunderscore}semi{\isacharunderscore}axioms\ sum{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isachardot}comm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymoplus}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}y\ {\isasymoplus}\ {\isacharquery}x{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isachardot}lcomm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymoplus}\ {\isacharparenleft}{\isacharquery}y\ {\isasymoplus}\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ {\isacharquery}y\ {\isasymoplus}\ {\isacharparenleft}{\isacharquery}x\ {\isasymoplus}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} \\ - \textbf{fixes} & \isa{prod} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}b{\isacharcomma}\ {\isacharprime}b{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequote}} - ~(\textbf{infixl}~\isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}}~70) \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{prod{\isachardot}assoc}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z\ {\isacharequal}\ {\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} \\ - \textbf{assumes} & \isa{{\isachardoublequote}comm{\isacharunderscore}semi{\isacharunderscore}axioms\ prod{\isachardoublequote}} \\ - \textbf{notes} & \isa{prod{\isachardot}comm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharquery}x{\isachardoublequote}} \\ - \textbf{notes} & \isa{prod{\isachardot}lcomm}: \isa{{\isachardoublequote}{\isacharquery}x\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}y\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ {\isacharquery}y\ {\isasymcdot}\ {\isacharparenleft}{\isacharquery}x\ {\isasymcdot}\ {\isacharquery}z{\isacharparenright}{\isachardoublequote}} \\ - \textbf{fixes} & \isa{hom} :: \isa{{\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequote}} \\ - \textbf{assumes} & \isa{{\isachardoublequote}semi{\isacharunderscore}hom{\isacharunderscore}axioms\ sum{\isachardoublequote}} \\ - \textbf{notes} & \isa{sum{\isacharunderscore}prod{\isacharunderscore}hom{\isachardot}hom}: - \isa{hom\ {\isacharparenleft}x\ {\isasymoplus}\ y{\isacharparenright}\ {\isacharequal}\ hom\ x\ {\isasymcdot}\ hom\ y} -\end{array} -\] - -\paragraph{Example 5.} - In this example, a locale expression leading to a list of context - elements that is not well-defined is encountered, and it is illustrated - how normalisation deals with multiple inheritance. - Consider the specification of monads (in the algebraic sense) - and monoids.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ monad\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ prod\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymcdot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{and}\ one\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ l{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymcdot}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ r{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymcdot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Monoids are both semigroups and monads and one would want to - specify them as locale expression \isa{semi\ {\isacharplus}\ monad}. - Unfortunately, this expression is not well-formed. Its normal form -\begin{align*% -} - \N(\isa{monad}) & = [\isa{monad\ prod}] \\ - \N(\isa{semi}+\isa{monad}) & = - \N(\isa{semi}) \App \N(\isa{monad}) - = [\isa{semi\ prod}, \isa{monad\ prod}] -\end{align*% -} - leads to a list containing the context element -\[ - \textbf{fixes}~\isa{prod} :: \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}} - ~(\textbf{infixl}~\isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}}~70) -\] - twice and thus violating the first criterion of well-formedness. To - avoid this problem, one can - introduce a new locale \isa{magma} with the sole purpose of fixing the - parameter and defining its syntax. The specifications of semigroup - and monad are changed so that they import \isa{magma}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ magma\ {\isacharequal}\ \isakeyword{fixes}\ prod\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymcdot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\isanewline -\isacommand{locale}\isamarkupfalse% -\ semi{\isacharprime}\ {\isacharequal}\ magma\ {\isacharplus}\ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymcdot}\ y{\isacharparenright}\ {\isasymcdot}\ z\ {\isacharequal}\ x\ {\isasymcdot}\ {\isacharparenleft}y\ {\isasymcdot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isacommand{locale}\isamarkupfalse% -\ monad{\isacharprime}\ {\isacharequal}\ magma\ {\isacharplus}\ \isakeyword{fixes}\ one\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ l{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymcdot}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ r{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymcdot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Normalisation now yields -\begin{align*% -} - \N(\isa{semi{\isacharprime}\ {\isacharplus}\ monad{\isacharprime}}) & = - \N(\isa{semi{\isacharprime}}) \App \N(\isa{monad{\isacharprime}}) \\ - & = (\N(\isa{magma}) \App [\isa{semi{\isacharprime}\ prod}]) \App - (\N(\isa{magma}) \App [\isa{monad{\isacharprime}\ prod}]) \\ - & = [\isa{magma\ prod}, \isa{semi{\isacharprime}\ prod}] \App - [\isa{magma\ prod}, \isa{monad{\isacharprime}\ prod}]) \\ - & = [\isa{magma\ prod}, \isa{semi{\isacharprime}\ prod}, - \isa{monad{\isacharprime}\ prod}] -\end{align*% -} - where the second occurrence of \isa{magma\ prod} is eliminated. - The reader is encouraged to check, using the \textbf{print\_locale} - command, that the list of context elements generated from this is - indeed well-formed. - - It follows that importing - parameters is more flexible than fixing them using a context element. - The Locale package provides the predefined locale \isa{var} that - can be used to import parameters if no - particular mixfix syntax is required. - Its definition is -\begin{center} - \textbf{locale} \isa{var} = \textbf{fixes} \isa{x{\isacharunderscore}} -\end{center} - The use of the internal variable \isa{x{\isacharunderscore}} - enforces that the parameter is renamed before being used, because - internal variables may not occur in the input syntax. Using - \isa{var}, the locale \isa{magma} could have been replaced by - the locale expression -\begin{center} - \isa{var} \isa{prod} (\textbf{infixl} \isa{{\isachardoublequote}{\isasymcdot}{\isachardoublequote}} 70) -\end{center} - in the above locale declarations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Includes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{sec-includes} - The context element \textbf{includes} takes a locale expression $e$ - as argument. It can only occur in long goals, where it - adds, like a target, locale context to the proof context. Unlike - with targets, the proved theorem is not stored - in the locale. Instead, it is exported immediately.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ lcomm{\isadigit{2}}{\isacharcolon}\isanewline -\ \ \isakeyword{includes}\ comm{\isacharunderscore}semi\ \isakeyword{shows}\ {\isachardoublequoteopen}x\ {\isasymcdot}\ {\isacharparenleft}y\ {\isasymcdot}\ z{\isacharparenright}\ {\isacharequal}\ y\ {\isasymcdot}\ {\isacharparenleft}x\ {\isasymcdot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymcdot}\ {\isacharparenleft}y\ {\isasymcdot}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymcdot}\ y{\isacharparenright}\ {\isasymcdot}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ {\isacharparenleft}y\ {\isasymcdot}\ x{\isacharparenright}\ {\isasymcdot}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ comm{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ y\ {\isasymcdot}\ {\isacharparenleft}x\ {\isasymcdot}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -This proof is identical to the proof of \isa{lcomm}. The use of - \textbf{includes} provides the same context and facts as when using - \isa{comm{\isacharunderscore}semi} as target. On the other hand, \isa{lcomm{\isadigit{2}}} is not added as a fact to the locale \isa{comm{\isacharunderscore}semi}, but - is directly visible in the theory. The theorem is -\[ - \isa{comm{\isacharunderscore}semi\ {\isacharquery}prod\ {\isasymLongrightarrow}\ {\isacharquery}prod\ {\isacharquery}x\ {\isacharparenleft}{\isacharquery}prod\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ {\isacharquery}prod\ {\isacharquery}y\ {\isacharparenleft}{\isacharquery}prod\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}}. -\] - Note that it is possible to - combine a target and (several) \textbf{includes} in a goal statement, thus - using contexts of several locales but storing the theorem in only - one of them.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Structures% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{sec-structures} - The specifications of semigroups and monoids that served as examples - in previous sections modelled each operation of an algebraic - structure as a single parameter. This is rather inconvenient for - structures with many operations, and also unnatural. In accordance - to mathematical texts, one would rather fix two groups instead of - two sets of operations. - - The approach taken in Isabelle is to encode algebraic structures - with suitable types (in Isabelle/HOL usually records). An issue to - be addressed by - locales is syntax for algebraic structures. This is the purpose of - the \textbf{(structure)} annotation in \textbf{fixes}, introduced by - Wenzel. We illustrate this, independently of record types, with a - different formalisation of semigroups. - - Let \isa{{\isacharprime}a\ semi{\isacharunderscore}type} be a not further specified type that - represents semigroups over the carrier type \isa{{\isacharprime}a}. Let \isa{s{\isacharunderscore}op} be an operation that maps an object of \isa{{\isacharprime}a\ semi{\isacharunderscore}type} to - a binary operation.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\isamarkupfalse% -\ {\isacharprime}a\ semi{\isacharunderscore}type\isanewline -\isacommand{consts}\isamarkupfalse% -\ s{\isacharunderscore}op\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharprime}a\ semi{\isacharunderscore}type{\isacharcomma}\ {\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymstar}{\isasymindex}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}% -\begin{isamarkuptext}% -Although \isa{s{\isacharunderscore}op} is a ternary operation, it is declared - infix. The syntax annotation contains the token \isa{{\isasymindex}} - (\verb.\.), which refers to the first argument. This syntax is only - effective in the context of a locale, and only if the first argument - is a - \emph{structural} parameter --- that is, a parameter with annotation - \textbf{(structure)}. The token has the effect of subscripting the - parameter --- by bracketing it between \verb.\<^bsub>. and \verb.\<^esub>.. - Additionally, the subscript of the first structural parameter may be - omitted, as in this specification of semigroups with structures:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ comm{\isacharunderscore}semi{\isacharprime}\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ G\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ semi{\isacharunderscore}type{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{structure}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymstar}\ y{\isacharparenright}\ {\isasymstar}\ z\ {\isacharequal}\ x\ {\isasymstar}\ {\isacharparenleft}y\ {\isasymstar}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{and}\ comm{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymstar}\ y\ {\isacharequal}\ y\ {\isasymstar}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Here \isa{x\ {\isasymstar}\ y} is equivalent to \isa{x\ {\isasymstar}\isactrlbsub G\isactrlesub \ y} and - abbreviates \isa{s{\isacharunderscore}op\ G\ x\ y}. A specification of homomorphisms - requires a second structural parameter.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ semi{\isacharprime}{\isacharunderscore}hom\ {\isacharequal}\ comm{\isacharunderscore}semi{\isacharprime}\ {\isacharplus}\ comm{\isacharunderscore}semi{\isacharprime}\ H\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ hom\isanewline -\ \ \isakeyword{assumes}\ hom{\isacharcolon}\ {\isachardoublequoteopen}hom\ {\isacharparenleft}x\ {\isasymstar}\ y{\isacharparenright}\ {\isacharequal}\ hom\ x\ {\isasymstar}\isactrlbsub H\isactrlesub \ hom\ y{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The parameter \isa{H} is defined in the second \textbf{fixes} - element of $\C(\isa{semi{\isacharprime}{\isacharunderscore}comm})$. Hence \isa{{\isasymstar}\isactrlbsub H\isactrlesub } - abbreviates \isa{s{\isacharunderscore}op\ H\ x\ y}. The same construction can be done - with records instead of an \textit{ad-hoc} type.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{record}\isamarkupfalse% -\ {\isacharprime}a\ semi\ {\isacharequal}\ prod\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}a{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymbullet}{\isasymindex}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}% -\begin{isamarkuptext}% -This declares the types \isa{{\isacharprime}a\ semi} and \isa{{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ semi{\isacharunderscore}scheme}. The latter is an extensible record, where the second - type argument is the type of the extension field. For details on - records, see \cite{NipkowEtAl2002} Chapter~8.3.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ semi{\isacharunderscore}w{\isacharunderscore}records\ {\isacharequal}\ struct\ G\ {\isacharplus}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymbullet}\ y{\isacharparenright}\ {\isasymbullet}\ z\ {\isacharequal}\ x\ {\isasymbullet}\ {\isacharparenleft}y\ {\isasymbullet}\ z{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The type \isa{{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ semi{\isacharunderscore}scheme} is inferred for the parameter \isa{G}. Using subtyping on records, the specification can be extended - to groups easily.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{record}\isamarkupfalse% -\ {\isacharprime}a\ group\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharprime}a\ semi{\isachardoublequoteclose}\ {\isacharplus}\isanewline -\ \ one\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}l{\isasymindex}{\isachardoublequoteclose}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ inv\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}inv{\isasymindex}\ {\isacharunderscore}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{8}}{\isadigit{1}}{\isacharbrackright}\ {\isadigit{8}}{\isadigit{0}}{\isacharparenright}\isanewline -\isacommand{locale}\isamarkupfalse% -\ group{\isacharunderscore}w{\isacharunderscore}records\ {\isacharequal}\ semi{\isacharunderscore}w{\isacharunderscore}records\ {\isacharplus}\isanewline -\ \ \isakeyword{assumes}\ l{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequoteopen}l\ {\isasymbullet}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ l{\isacharunderscore}inv{\isacharcolon}\ {\isachardoublequoteopen}inv\ x\ {\isasymbullet}\ x\ {\isacharequal}\ l{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Finally, the predefined locale -\begin{center} - \textbf{locale} \textit{struct} = \textbf{fixes} \isa{S{\isacharunderscore}} - \textbf{(structure)}. -\end{center} - is analogous to \isa{var}. - More examples on the use of structures, including groups, rings and - polynomials can be found in the Isabelle distribution in the - session HOL-Algebra.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Conclusions and Outlook% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locales provide a simple means of modular reasoning. They enable to - abbreviate frequently occurring context statements and maintain facts - valid in these contexts. Importantly, using structures, they allow syntax to be - effective only in certain contexts, and thus to mimic common - practice in mathematics, where notation is chosen very flexibly. - This is also known as literate formalisation \cite{Bailey1998}. - Locale expressions allow to duplicate and merge - specifications. This is a necessity, for example, when reasoning about - homomorphisms. Normalisation makes it possible to deal with - diamond-shaped inheritance structures, and generally with directed - acyclic graphs. The combination of locales with record - types in higher-order logic provides an effective means for - specifying algebraic structures: locale import and record subtyping - provide independent hierarchies for specifications and structure - elements. Rich examples for this can be found in - the Isabelle distribution in the session HOL-Algebra. - - The primary reason for writing this report was to provide a better - understanding of locales in Isar. Wenzel provided hardly any - documentation, with the exception of \cite{Wenzel2002b}. The - present report should make it easier for users of Isabelle to take - advantage of locales. - - The report is also a base for future extensions. These include - improved syntax for structures. Identifying them by numbers seems - unnatural and can be confusing if more than two structures are - involved --- for example, when reasoning about universal - properties --- and numbering them by order of occurrence seems - arbitrary. Another desirable feature is \emph{instantiation}. One - may, in the course of a theory development, construct objects that - fulfil the specification of a locale. These objects are possibly - defined in the context of another locale. Instantiation should make it - simple to specialise abstract facts for the object under - consideration and to use the specified facts. - - A detailed comparison of locales with module systems in type theory - has not been undertaken yet, but could be beneficial. For example, - a module system for Coq has recently been presented by Chrzaszcz - \cite{Chrzaszcz2003,Chrzaszcz2004}. While the - latter usually constitute extensions of the calculus, locales are - a rather thin layer that does not change Isabelle's meta logic. - Locales mainly manage specifications and facts. Functors, like - the constructor for polynomial rings, remain objects of the - logic. - - \textbf{Acknowledgements.} Lawrence C.\ Paulson and Norbert - Schirmer made useful comments on a draft of this paper.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/root.bib --- a/doc-src/Locales/Locales/document/root.bib Tue Jun 03 11:55:35 2008 +0200 +++ b/doc-src/Locales/Locales/document/root.bib Tue Jun 03 12:34:22 2008 +0200 @@ -1,173 +1,81 @@ -@phdthesis{Bailey1998, - author = "Anthony Bailey", - title = "The machine-checked literate formalisation of algebra in type theory", - school = "University of Manchester", - month = jan, - year = 1998, - available = { CB } +@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}." } -@phdthesis{Kammuller1999a, - author = "Florian Kamm{\"u}ller", - title = "Modular Reasoning in {Isabelle}", - school = "University of Cambridge, Computer Laboratory", - note = "Available as Technical Report No. 470.", - month = aug, - year = 1999, - available = { CB } -} - -% TYPES 98 - -@inproceedings{Kammuller1999b, - author = "Florian Kamm{\"u}ller", - title = "Modular Structures as Dependent Types in {Isabelle}", - pages = "121--132", - crossref = "AltenkirchEtAl1999", +@book {Jacobson1985, + author = "Nathan Jacobson", + title = "Basic Algebra", + volume = "I", + publisher = "Freeman", + edition = "2nd", + year = 1985, available = { CB } } -@proceedings{AltenkirchEtAl1999, - editor = "Thorsten Altenkirch and Wolfgang Naraschewski and Bernhard Reus", - title = "Types for Proofs and Programs: International Workshop, {TYPES} '98, {Kloster Irsee, Germany, March 27--31, 1998}, Selected Papers", - booktitle = "TYPES '98", - publisher = "Springer", - series = "LNCS 1657", - year = 1999 -} -% CADE-17 +% TYPES 2006 -@inproceedings{Kammuller2000, - author = "Florian Kamm{\"u}ller", - title = "Modular Reasoning in {Isabelle}", - pages = "99--114", - crossref = "McAllester2000", - available = { CB } -} - -@proceedings{McAllester2000, - editor = "David McAllester", - title = "17th International Conference on Automated Deduction, Pittsburgh, PA, USA, June 17--20, 2000: Proceedings", - booktitle = "CADE 17", - publisher = "Springer", - series = "LNCS 1831", - year = 2000 -} - -@book{NipkowEtAl2002, - author = "Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel", - title = "{Isabelle/HOL}: A Proof Assistant for Higher-Order Logic", - publisher = "Springer", - series = "LNCS 2283", - year = 2002, +@inproceedings{HaftmannWenzel2007, + author = "Florian Haftmann and Makarius Wenzel", + title = "Constructive Type Classes in {Isabelle}", + pages = "160--174", + crossref = "AltenkirchMcBride2007", available = { CB } } -% TYPES 2002 - -@inproceedings{Nipkow2003, - author = "Tobias Nipkow", - title = "Structured Proofs in {Isar/HOL}", - pages = "259--278", - crossref = "GeuversWiedijk2003" -} - -@proceedings{GeuversWiedijk2003, - editor = "H. Geuvers and F. Wiedijk", - title = "Types for Proofs and Programs (TYPES 2002)", - booktitle = "TYPES 2002", +@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 2646", - year = 2003 + series = "LNCS 4502", + year = 2007 } -@phdthesis{Wenzel2002a, - author = "Markus Wenzel", - title = "{Isabelle/Isar} --- a versatile environment for human-readable formal proof documents", - school = "Technische Universit{\"a}t M{\"u}nchen", - note = "Electronically published as http://tumb1.biblio.tu-muenchen.de/publ/diss/in/2002/wenzel.html.", - year = 2002 -} -@unpublished{Wenzel2002b, - author = "Markus Wenzel", - title = "Using locales in {Isabelle/Isar}", - note = "Part of the Isabelle2003 distribution, file src/HOL/ex/Locales.thy. Distribution of Isabelle available at http://isabelle.in.tum.de", - year = 2002 +@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 } -@unpublished{Wenzel2003, - author = "Markus Wenzel", - title = "The {Isabelle/Isar} Reference Manual", - note = "Part of the Isabelle2003 distribution, available at http://isabelle.in.tum.de", - year = 2003 +% MKM 2006 + +@inproceedings{Ballarin2006b, + author = "Clemens Ballarin", + title = "Interpretation of Locales in {Isabelle}: Theories and Proof Contexts", + pages = "31--43", + crossref = "BorweinFarmer2006" } -% TPHOLs 2003 - -@inproceedings{Chrzaszcz2003, - author = "Jacek Chrz{\c{a}}szcz", - title = "Implementing Modules in the {Coq} System", - pages = "270--286", - crossref = "BasinWolff2003", +@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 } } -@proceedings{BasinWolff2003, - editor = "David Basin and Burkhart Wolff", - title = "Theorem proving in higher order logics: 16th International Conference, TPHOLs 2003, Rome, Italy, September 2003: proceedings", - booktitle = "TPHOLs 2003", - publisher = "Springer", - series = "LNCS 2758", - year = 2003 -} +% TPHOLs 1999 -@PhdThesis{Klein2003, - author = "Gerwin Klein", - title = "Verified Java Bytecode Verification", - school = "Institut f{\"u}r Informatik, Technische Universit{\"a}t M{\"u}nchen", - year = 2003 -} - -% TACAS 2000 - -@inproceedings{Aspinall2000, - author = "David Aspinall", - title = "Proof General: A Generic Tool for Proof Development", - pages = "38--42", - crossref = "GrafSchwartzbach2000" +@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 } } -@proceedings{GrafSchwartzbach2000, - editor = {Susanne Graf and Michael I. Schwartzbach}, - title = {Tools and Algorithms for Construction and Analysis of Systems, 6th International Conference, TACAS 2000, Berlin, Germany, March 25 -- April 2, 2000, Proceedings}, - booktitle = "TACAS 2000", - publisher = {Springer}, - series = {LNCS 1785}, - year = {2000}, -} - -% TYPES 2004 - -@inproceedings{Ballarin2004a, - author = "Clemens Ballarin", - title = "Locales and Locale Expressions in {Isabelle/Isar}", - pages = "34--50", - crossref = "BerardiEtAl2004" +@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 } - -@inproceedings{Chrzaszcz2004, - author = "Jacek Chrz{\c{a}}szcz", - title = "Modules in {Coq} Are and Will Be Correct", - pages = "130--136", - crossref = "BerardiEtAl2004", - available = { CB } -} -@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 -} diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/root.tex --- a/doc-src/Locales/Locales/document/root.tex Tue Jun 03 11:55:35 2008 +0200 +++ b/doc-src/Locales/Locales/document/root.tex Tue Jun 03 12:34:22 2008 +0200 @@ -1,48 +1,81 @@ -\documentclass[leqno]{article} -\usepackage{../isabelle,../isabellesym} - +\documentclass[11pt,a4paper]{article} \usepackage{amsmath} -\usepackage{amssymb} +\usepackage{tikz} +\usepackage{subfigure} +\usepackage{isabelle,isabellesym} +\usepackage{verbatim} \usepackage{array} +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage[greek,english]{babel} + %option greek for \ + %option english (default language) for \, \ + +%\usepackage[latin1]{inputenc} + %for \, \, \, \, + %\, \, \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \ + % this should be the last package used -\usepackage{../pdfsetup} +\usepackage{pdfsetup} +\isafoldtag{proof} + +% urls in roman style, theory text in typewriter \urlstyle{rm} \isabellestyle{tt} -%\renewcommand{\isacharunderscore}{\_}% -% the latter is not necessary with \isabellestyle{tt} + \begin{document} -\title{Locales and Locale Expressions in Isabelle/Isar} -\author{Clemens Ballarin \\ - Fakult\"at f\"ur Informatik \\ Technische Universit\"at M\"unchen \\ - 85748 Garching, Germany \\ - ballarin@in.tum.de} +\title{Tutorial to Locales and Locale Interpretation} +\author{Clemens Ballarin} \date{} \maketitle +\thispagestyle{myheadings} +\markright{Technical Report TUM-I0723, Technische Universit\"at M\"unchen, 2007} + \begin{abstract} - Locales provide a module system for the Isabelle proof assistant. - Recently, locales have been ported to the new Isar format for - structured proofs. At the same time, they have been extended by - locale expressions, a language for composing locale specifications, - and by structures, which provide syntax for algebraic structures. - The present paper presents both and is suitable as a tutorial to - locales in Isar, because it covers both basics and recent - extensions, and contains many examples. + Locales are Isabelle's mechanism to deal with parametric theories. + We present typical examples of locale specifications, + along with interpretations between locales to change their + hierarchic dependencies and interpretations to reuse locales in + theory contexts and proofs. + + This tutorial is intended for locale novices; familiarity with + Isabelle and Isar is presumed. \end{abstract} -%\tableofcontents - \parindent 0pt\parskip 0.5ex -% include generated text of all theories +% generated text of all theories \input{session} -\bibliographystyle{plain} +% optional bibliography +\newpage +\bibliographystyle{abbrv} \bibliography{root} \end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Locales/document/session.tex --- a/doc-src/Locales/Locales/document/session.tex Tue Jun 03 11:55:35 2008 +0200 +++ b/doc-src/Locales/Locales/document/session.tex Tue Jun 03 12:34:22 2008 +0200 @@ -1,4 +1,12 @@ -\input{Locales.tex} +\input{GCD.tex} + +\input{Examples.tex} + +\input{Examples1.tex} + +\input{Examples2.tex} + +\input{Examples3.tex} %%% Local Variables: %%% mode: latex diff -r 9681f347b6f5 -r d1d35284542f doc-src/Locales/Makefile --- a/doc-src/Locales/Makefile Tue Jun 03 11:55:35 2008 +0200 +++ b/doc-src/Locales/Makefile Tue Jun 03 12:34:22 2008 +0200 @@ -17,7 +17,9 @@ NAME = locales FILES = Locales/document/root.tex Locales/document/root.bib \ - Locales/document/session.tex Locales/document/Locales.tex \ + Locales/document/session.tex Locales/document/Examples.tex \ + Locales/document/Examples1.tex Locales/document/Examples2.tex \ + Locales/document/Examples3.tex \ ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty dvi: $(NAME).dvi