# HG changeset patch # User haftmann # Date 1236074451 -3600 # Node ID 2f4684e2ea954d317d5c0f300ad35c6d38ae6826 # Parent 2775062fd3a9624dddba06dd2f821dc32c60ad6a more canonical directory structure of manuals diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/IsaMakefile Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,33 @@ + +## targets + +default: Thy +images: +test: Thy + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document + + +## Thy + +THY = $(LOG)/HOL-Thy.gz + +Thy: $(THY) + +$(THY): Thy/ROOT.ML Thy/Setup.thy Thy/Classes.thy ../antiquote_setup.ML ../more_antiquote.ML + @$(USEDIR) HOL Thy + + +## clean + +clean: + @rm -f $(THY) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Makefile Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,35 @@ + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = classes + +FILES = $(NAME).tex classes.tex Thy/document/Classes.tex \ + style.sty ../iman.sty ../extra.sty ../isar.sty \ + ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ + ../manual.bib ../proof.sty + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_isar.eps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_isar.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/Thy/Classes.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/Classes.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,635 @@ +theory Classes +imports Main Setup +begin + +chapter {* Haskell-style classes with Isabelle/Isar *} + +section {* Introduction *} + +text {* + Type classes were introduces by Wadler and Blott \cite{wadler89how} + into the Haskell language, to allow for a reasonable implementation + of overloading\footnote{throughout this tutorial, we are referring + to classical Haskell 1.0 type classes, not considering + later additions in expressiveness}. + As a canonical example, a polymorphic equality function + @{text "eq \ \ \ \ \ bool"} which is overloaded on different + types for @{text "\"}, which is achieved by splitting introduction + of the @{text eq} function from its overloaded definitions by means + of @{text class} and @{text instance} declarations: + + \begin{quote} + + \noindent@{text "class eq where"}\footnote{syntax here is a kind of isabellized Haskell} \\ + \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} + + \medskip\noindent@{text "instance nat \ eq where"} \\ + \hspace*{2ex}@{text "eq 0 0 = True"} \\ + \hspace*{2ex}@{text "eq 0 _ = False"} \\ + \hspace*{2ex}@{text "eq _ 0 = False"} \\ + \hspace*{2ex}@{text "eq (Suc n) (Suc m) = eq n m"} + + \medskip\noindent@{text "instance (\\eq, \\eq) pair \ eq where"} \\ + \hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2"} + + \medskip\noindent@{text "class ord extends eq where"} \\ + \hspace*{2ex}@{text "less_eq \ \ \ \ \ bool"} \\ + \hspace*{2ex}@{text "less \ \ \ \ \ bool"} + + \end{quote} + + \noindent Type variables are annotated with (finitely many) classes; + these annotations are assertions that a particular polymorphic type + provides definitions for overloaded functions. + + Indeed, type classes not only allow for simple overloading + but form a generic calculus, an instance of order-sorted + algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. + + From a software engeneering point of view, type classes + roughly correspond to interfaces in object-oriented languages like Java; + so, it is naturally desirable that type classes do not only + provide functions (class parameters) but also state specifications + implementations must obey. For example, the @{text "class eq"} + above could be given the following specification, demanding that + @{text "class eq"} is an equivalence relation obeying reflexivity, + symmetry and transitivity: + + \begin{quote} + + \noindent@{text "class eq where"} \\ + \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \\ + @{text "satisfying"} \\ + \hspace*{2ex}@{text "refl: eq x x"} \\ + \hspace*{2ex}@{text "sym: eq x y \ eq x y"} \\ + \hspace*{2ex}@{text "trans: eq x y \ eq y z \ eq x z"} + + \end{quote} + + \noindent From a theoretic point of view, type classes are lightweight + modules; Haskell type classes may be emulated by + SML functors \cite{classes_modules}. + Isabelle/Isar offers a discipline of type classes which brings + all those aspects together: + + \begin{enumerate} + \item specifying abstract parameters together with + corresponding specifications, + \item instantiating those abstract parameters by a particular + type + \item in connection with a ``less ad-hoc'' approach to overloading, + \item with a direct link to the Isabelle module system + (aka locales \cite{kammueller-locales}). + \end{enumerate} + + \noindent Isar type classes also directly support code generation + in a Haskell like fashion. + + This tutorial demonstrates common elements of structured specifications + and abstract reasoning with type classes by the algebraic hierarchy of + semigroups, monoids and groups. Our background theory is that of + Isabelle/HOL \cite{isa-tutorial}, for which some + familiarity is assumed. + + Here we merely present the look-and-feel for end users. + Internally, those are mapped to more primitive Isabelle concepts. + See \cite{Haftmann-Wenzel:2006:classes} for more detail. +*} + +section {* A simple algebra example \label{sec:example} *} + +subsection {* Class definition *} + +text {* + Depending on an arbitrary type @{text "\"}, class @{text + "semigroup"} introduces a binary operator @{text "(\)"} that is + assumed to be associative: +*} + +class %quote semigroup = + fixes mult :: "\ \ \ \ \" (infixl "\" 70) + assumes assoc: "(x \ y) \ z = x \ (y \ z)" + +text {* + \noindent This @{command class} specification consists of two + parts: the \qn{operational} part names the class parameter + (@{element "fixes"}), the \qn{logical} part specifies properties on them + (@{element "assumes"}). The local @{element "fixes"} and + @{element "assumes"} are lifted to the theory toplevel, + yielding the global + parameter @{term [source] "mult \ \\semigroup \ \ \ \"} and the + global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y + z \ \\semigroup. (x \ y) \ z = x \ (y \ z)"}. +*} + + +subsection {* Class instantiation \label{sec:class_inst} *} + +text {* + The concrete type @{typ int} is made a @{class semigroup} + instance by providing a suitable definition for the class parameter + @{text "(\)"} and a proof for the specification of @{fact assoc}. + This is accomplished by the @{command instantiation} target: +*} + +instantiation %quote int :: semigroup +begin + +definition %quote + mult_int_def: "i \ j = i + (j\int)" + +instance %quote proof + fix i j k :: int have "(i + j) + k = i + (j + k)" by simp + then show "(i \ j) \ k = i \ (j \ k)" + unfolding mult_int_def . +qed + +end %quote + +text {* + \noindent @{command instantiation} allows to define class parameters + at a particular instance using common specification tools (here, + @{command definition}). The concluding @{command instance} + opens a proof that the given parameters actually conform + to the class specification. Note that the first proof step + is the @{method default} method, + which for such instance proofs maps to the @{method intro_classes} method. + This boils down an instance judgement to the relevant primitive + proof goals and should conveniently always be the first method applied + in an instantiation proof. + + From now on, the type-checker will consider @{typ int} + as a @{class semigroup} automatically, i.e.\ any general results + are immediately available on concrete instances. + + \medskip Another instance of @{class semigroup} are the natural numbers: +*} + +instantiation %quote nat :: semigroup +begin + +primrec %quote mult_nat where + "(0\nat) \ n = n" + | "Suc m \ n = Suc (m \ n)" + +instance %quote proof + fix m n q :: nat + show "m \ n \ q = m \ (n \ q)" + by (induct m) auto +qed + +end %quote + +text {* + \noindent Note the occurence of the name @{text mult_nat} + in the primrec declaration; by default, the local name of + a class operation @{text f} to instantiate on type constructor + @{text \} are mangled as @{text f_\}. In case of uncertainty, + these names may be inspected using the @{command "print_context"} command + or the corresponding ProofGeneral button. +*} + +subsection {* Lifting and parametric types *} + +text {* + Overloaded definitions giving on class instantiation + may include recursion over the syntactic structure of types. + As a canonical example, we model product semigroups + using our simple algebra: +*} + +instantiation %quote * :: (semigroup, semigroup) semigroup +begin + +definition %quote + mult_prod_def: "p\<^isub>1 \ p\<^isub>2 = (fst p\<^isub>1 \ fst p\<^isub>2, snd p\<^isub>1 \ snd p\<^isub>2)" + +instance %quote proof + fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\\semigroup \ \\semigroup" + show "p\<^isub>1 \ p\<^isub>2 \ p\<^isub>3 = p\<^isub>1 \ (p\<^isub>2 \ p\<^isub>3)" + unfolding mult_prod_def by (simp add: assoc) +qed + +end %quote + +text {* + \noindent Associativity from product semigroups is + established using + the definition of @{text "(\)"} on products and the hypothetical + associativity of the type components; these hypotheses + are facts due to the @{class semigroup} constraints imposed + on the type components by the @{command instance} proposition. + Indeed, this pattern often occurs with parametric types + and type classes. +*} + + +subsection {* Subclassing *} + +text {* + We define a subclass @{text monoidl} (a semigroup with a left-hand neutral) + by extending @{class semigroup} + with one additional parameter @{text neutral} together + with its property: +*} + +class %quote monoidl = semigroup + + fixes neutral :: "\" ("\") + assumes neutl: "\ \ x = x" + +text {* + \noindent Again, we prove some instances, by + providing suitable parameter definitions and proofs for the + additional specifications. Observe that instantiations + for types with the same arity may be simultaneous: +*} + +instantiation %quote nat and int :: monoidl +begin + +definition %quote + neutral_nat_def: "\ = (0\nat)" + +definition %quote + neutral_int_def: "\ = (0\int)" + +instance %quote proof + fix n :: nat + show "\ \ n = n" + unfolding neutral_nat_def by simp +next + fix k :: int + show "\ \ k = k" + unfolding neutral_int_def mult_int_def by simp +qed + +end %quote + +instantiation %quote * :: (monoidl, monoidl) monoidl +begin + +definition %quote + neutral_prod_def: "\ = (\, \)" + +instance %quote proof + fix p :: "\\monoidl \ \\monoidl" + show "\ \ p = p" + unfolding neutral_prod_def mult_prod_def by (simp add: neutl) +qed + +end %quote + +text {* + \noindent Fully-fledged monoids are modelled by another subclass + which does not add new parameters but tightens the specification: +*} + +class %quote monoid = monoidl + + assumes neutr: "x \ \ = x" + +instantiation %quote nat and int :: monoid +begin + +instance %quote proof + fix n :: nat + show "n \ \ = n" + unfolding neutral_nat_def by (induct n) simp_all +next + fix k :: int + show "k \ \ = k" + unfolding neutral_int_def mult_int_def by simp +qed + +end %quote + +instantiation %quote * :: (monoid, monoid) monoid +begin + +instance %quote proof + fix p :: "\\monoid \ \\monoid" + show "p \ \ = p" + unfolding neutral_prod_def mult_prod_def by (simp add: neutr) +qed + +end %quote + +text {* + \noindent To finish our small algebra example, we add a @{text group} class + with a corresponding instance: +*} + +class %quote group = monoidl + + fixes inverse :: "\ \ \" ("(_\
)" [1000] 999) + assumes invl: "x\
\ x = \" + +instantiation %quote int :: group +begin + +definition %quote + inverse_int_def: "i\
= - (i\int)" + +instance %quote proof + fix i :: int + have "-i + i = 0" by simp + then show "i\
\ i = \" + unfolding mult_int_def neutral_int_def inverse_int_def . +qed + +end %quote + + +section {* Type classes as locales *} + +subsection {* A look behind the scene *} + +text {* + The example above gives an impression how Isar type classes work + in practice. As stated in the introduction, classes also provide + a link to Isar's locale system. Indeed, the logical core of a class + is nothing else than a locale: +*} + +class %quote idem = + fixes f :: "\ \ \" + assumes idem: "f (f x) = f x" + +text {* + \noindent essentially introduces the locale +*} setup %invisible {* Sign.add_path "foo" *} + +locale %quote idem = + fixes f :: "\ \ \" + assumes idem: "f (f x) = f x" + +text {* \noindent together with corresponding constant(s): *} + +consts %quote f :: "\ \ \" + +text {* + \noindent The connection to the type system is done by means + of a primitive axclass +*} setup %invisible {* Sign.add_path "foo" *} + +axclass %quote idem < type + idem: "f (f x) = f x" setup %invisible {* Sign.parent_path *} + +text {* \noindent together with a corresponding interpretation: *} + +interpretation %quote idem_class: + idem "f \ (\\idem) \ \" +proof qed (rule idem) + +text {* + \noindent This gives you at hand the full power of the Isabelle module system; + conclusions in locale @{text idem} are implicitly propagated + to class @{text idem}. +*} setup %invisible {* Sign.parent_path *} + +subsection {* Abstract reasoning *} + +text {* + Isabelle locales enable reasoning at a general level, while results + are implicitly transferred to all instances. For example, we can + now establish the @{text "left_cancel"} lemma for groups, which + states that the function @{text "(x \)"} is injective: +*} + +lemma %quote (in group) left_cancel: "x \ y = x \ z \ y = z" +proof + assume "x \ y = x \ z" + then have "x\
\ (x \ y) = x\
\ (x \ z)" by simp + then have "(x\
\ x) \ y = (x\
\ x) \ z" using assoc by simp + then show "y = z" using neutl and invl by simp +next + assume "y = z" + then show "x \ y = x \ z" by simp +qed + +text {* + \noindent Here the \qt{@{keyword "in"} @{class group}} target specification + indicates that the result is recorded within that context for later + use. This local theorem is also lifted to the global one @{fact + "group.left_cancel:"} @{prop [source] "\x y z \ \\group. x \ y = x \ + z \ y = z"}. Since type @{text "int"} has been made an instance of + @{text "group"} before, we may refer to that fact as well: @{prop + [source] "\x y z \ int. x \ y = x \ z \ y = z"}. +*} + + +subsection {* Derived definitions *} + +text {* + Isabelle locales support a concept of local definitions + in locales: +*} + +primrec %quote (in monoid) pow_nat :: "nat \ \ \ \" where + "pow_nat 0 x = \" + | "pow_nat (Suc n) x = x \ pow_nat n x" + +text {* + \noindent If the locale @{text group} is also a class, this local + definition is propagated onto a global definition of + @{term [source] "pow_nat \ nat \ \\monoid \ \\monoid"} + with corresponding theorems + + @{thm pow_nat.simps [no_vars]}. + + \noindent As you can see from this example, for local + definitions you may use any specification tool + which works together with locales (e.g. \cite{krauss2006}). +*} + + +subsection {* A functor analogy *} + +text {* + We introduced Isar classes by analogy to type classes + functional programming; if we reconsider this in the + context of what has been said about type classes and locales, + we can drive this analogy further by stating that type + classes essentially correspond to functors which have + a canonical interpretation as type classes. + Anyway, there is also the possibility of other interpretations. + For example, also @{text list}s form a monoid with + @{text append} and @{term "[]"} as operations, but it + seems inappropriate to apply to lists + the same operations as for genuinely algebraic types. + In such a case, we simply can do a particular interpretation + of monoids for lists: +*} + +interpretation %quote list_monoid!: monoid append "[]" + proof qed auto + +text {* + \noindent This enables us to apply facts on monoids + to lists, e.g. @{thm list_monoid.neutl [no_vars]}. + + When using this interpretation pattern, it may also + be appropriate to map derived definitions accordingly: +*} + +primrec %quote replicate :: "nat \ \ list \ \ list" where + "replicate 0 _ = []" + | "replicate (Suc n) xs = xs @ replicate n xs" + +interpretation %quote list_monoid!: monoid append "[]" where + "monoid.pow_nat append [] = replicate" +proof - + interpret monoid append "[]" .. + show "monoid.pow_nat append [] = replicate" + proof + fix n + show "monoid.pow_nat append [] n = replicate n" + by (induct n) auto + qed +qed intro_locales + + +subsection {* Additional subclass relations *} + +text {* + Any @{text "group"} is also a @{text "monoid"}; this + can be made explicit by claiming an additional + subclass relation, + together with a proof of the logical difference: +*} + +subclass %quote (in group) monoid +proof + fix x + from invl have "x\
\ x = \" by simp + with assoc [symmetric] neutl invl have "x\
\ (x \ \) = x\
\ x" by simp + with left_cancel show "x \ \ = x" by simp +qed + +text {* + \noindent The logical proof is carried out on the locale level. + Afterwards it is propagated + to the type system, making @{text group} an instance of + @{text monoid} by adding an additional edge + to the graph of subclass relations + (cf.\ \figref{fig:subclass}). + + \begin{figure}[htbp] + \begin{center} + \small + \unitlength 0.6mm + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){@{text semigroup}}} + \put(20,40){\makebox(0,0){@{text monoidl}}} + \put(00,20){\makebox(0,0){@{text monoid}}} + \put(40,00){\makebox(0,0){@{text group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(25,35){\vector(1,-3){10}} + \end{picture} + \hspace{8em} + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){@{text semigroup}}} + \put(20,40){\makebox(0,0){@{text monoidl}}} + \put(00,20){\makebox(0,0){@{text monoid}}} + \put(40,00){\makebox(0,0){@{text group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(05,15){\vector(3,-1){30}} + \end{picture} + \caption{Subclass relationship of monoids and groups: + before and after establishing the relationship + @{text "group \ monoid"}; transitive edges are left out.} + \label{fig:subclass} + \end{center} + \end{figure} +7 + For illustration, a derived definition + in @{text group} which uses @{text pow_nat}: +*} + +definition %quote (in group) pow_int :: "int \ \ \ \" where + "pow_int k x = (if k >= 0 + then pow_nat (nat k) x + else (pow_nat (nat (- k)) x)\
)" + +text {* + \noindent yields the global definition of + @{term [source] "pow_int \ int \ \\group \ \\group"} + with the corresponding theorem @{thm pow_int_def [no_vars]}. +*} + +subsection {* A note on syntax *} + +text {* + As a commodity, class context syntax allows to refer + to local class operations and their global counterparts + uniformly; type inference resolves ambiguities. For example: +*} + +context %quote semigroup +begin + +term %quote "x \ y" -- {* example 1 *} +term %quote "(x\nat) \ y" -- {* example 2 *} + +end %quote + +term %quote "x \ y" -- {* example 3 *} + +text {* + \noindent Here in example 1, the term refers to the local class operation + @{text "mult [\]"}, whereas in example 2 the type constraint + enforces the global class operation @{text "mult [nat]"}. + In the global context in example 3, the reference is + to the polymorphic global class operation @{text "mult [?\ \ semigroup]"}. +*} + +section {* Further issues *} + +subsection {* Type classes and code generation *} + +text {* + Turning back to the first motivation for type classes, + namely overloading, it is obvious that overloading + stemming from @{command class} statements and + @{command instantiation} + targets naturally maps to Haskell type classes. + The code generator framework \cite{isabelle-codegen} + takes this into account. Concerning target languages + lacking type classes (e.g.~SML), type classes + are implemented by explicit dictionary construction. + As example, let's go back to the power function: +*} + +definition %quote example :: int where + "example = pow_int 10 (-2)" + +text {* + \noindent This maps to Haskell as: +*} + +text %quote {*@{code_stmts example (Haskell)}*} + +text {* + \noindent The whole code in SML with explicit dictionary passing: +*} + +text %quote {*@{code_stmts example (SML)}*} + +subsection {* Inspecting the type class universe *} + +text {* + To facilitate orientation in complex subclass structures, + two diagnostics commands are provided: + + \begin{description} + + \item[@{command "print_classes"}] print a list of all classes + together with associated operations etc. + + \item[@{command "class_deps"}] visualizes the subclass relation + between all classes as a Hasse diagram. + + \end{description} +*} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/ROOT.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,6 @@ + +(* $Id$ *) + +no_document use_thy "Setup"; + +use_thy "Classes"; diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/Thy/Setup.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/Setup.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,34 @@ +theory Setup +imports Main Code_Integer +uses + "../../../antiquote_setup" + "../../../more_antiquote" +begin + +ML {* Code_Target.code_width := 74 *} + +syntax + "_alpha" :: "type" ("\") + "_alpha_ofsort" :: "sort \ type" ("\()\_" [0] 1000) + "_beta" :: "type" ("\") + "_beta_ofsort" :: "sort \ type" ("\()\_" [0] 1000) + +parse_ast_translation {* + let + fun alpha_ast_tr [] = Syntax.Variable "'a" + | alpha_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); + fun alpha_ofsort_ast_tr [ast] = + Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'a", ast] + | alpha_ofsort_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); + fun beta_ast_tr [] = Syntax.Variable "'b" + | beta_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); + fun beta_ofsort_ast_tr [ast] = + Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'b", ast] + | beta_ofsort_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); + in [ + ("_alpha", alpha_ast_tr), ("_alpha_ofsort", alpha_ofsort_ast_tr), + ("_beta", beta_ast_tr), ("_beta_ofsort", beta_ofsort_ast_tr) + ] end +*} + +end \ No newline at end of file diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/Thy/document/Classes.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/document/Classes.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,1346 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Classes}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Classes\isanewline +\isakeyword{imports}\ Main\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Haskell-style classes with Isabelle/Isar% +} +\isamarkuptrue% +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Type classes were introduces by Wadler and Blott \cite{wadler89how} + into the Haskell language, to allow for a reasonable implementation + of overloading\footnote{throughout this tutorial, we are referring + to classical Haskell 1.0 type classes, not considering + later additions in expressiveness}. + As a canonical example, a polymorphic equality function + \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different + types for \isa{{\isasymalpha}}, which is achieved by splitting introduction + of the \isa{eq} function from its overloaded definitions by means + of \isa{class} and \isa{instance} declarations: + + \begin{quote} + + \noindent\isa{class\ eq\ where}\footnote{syntax here is a kind of isabellized Haskell} \\ + \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} + + \medskip\noindent\isa{instance\ nat\ {\isasymColon}\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isadigit{0}}\ {\isacharequal}\ True} \\ + \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ False} \\ + \hspace*{2ex}\isa{eq\ {\isacharunderscore}\ {\isadigit{0}}\ {\isacharequal}\ False} \\ + \hspace*{2ex}\isa{eq\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ eq\ n\ m} + + \medskip\noindent\isa{instance\ {\isacharparenleft}{\isasymalpha}{\isasymColon}eq{\isacharcomma}\ {\isasymbeta}{\isasymColon}eq{\isacharparenright}\ pair\ {\isasymColon}\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ eq\ x{\isadigit{1}}\ x{\isadigit{2}}\ {\isasymand}\ eq\ y{\isadigit{1}}\ y{\isadigit{2}}} + + \medskip\noindent\isa{class\ ord\ extends\ eq\ where} \\ + \hspace*{2ex}\isa{less{\isacharunderscore}eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ + \hspace*{2ex}\isa{less\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} + + \end{quote} + + \noindent Type variables are annotated with (finitely many) classes; + these annotations are assertions that a particular polymorphic type + provides definitions for overloaded functions. + + Indeed, type classes not only allow for simple overloading + but form a generic calculus, an instance of order-sorted + algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. + + From a software engeneering point of view, type classes + roughly correspond to interfaces in object-oriented languages like Java; + so, it is naturally desirable that type classes do not only + provide functions (class parameters) but also state specifications + implementations must obey. For example, the \isa{class\ eq} + above could be given the following specification, demanding that + \isa{class\ eq} is an equivalence relation obeying reflexivity, + symmetry and transitivity: + + \begin{quote} + + \noindent\isa{class\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ + \isa{satisfying} \\ + \hspace*{2ex}\isa{refl{\isacharcolon}\ eq\ x\ x} \\ + \hspace*{2ex}\isa{sym{\isacharcolon}\ eq\ x\ y\ {\isasymlongleftrightarrow}\ eq\ x\ y} \\ + \hspace*{2ex}\isa{trans{\isacharcolon}\ eq\ x\ y\ {\isasymand}\ eq\ y\ z\ {\isasymlongrightarrow}\ eq\ x\ z} + + \end{quote} + + \noindent From a theoretic point of view, type classes are lightweight + modules; Haskell type classes may be emulated by + SML functors \cite{classes_modules}. + Isabelle/Isar offers a discipline of type classes which brings + all those aspects together: + + \begin{enumerate} + \item specifying abstract parameters together with + corresponding specifications, + \item instantiating those abstract parameters by a particular + type + \item in connection with a ``less ad-hoc'' approach to overloading, + \item with a direct link to the Isabelle module system + (aka locales \cite{kammueller-locales}). + \end{enumerate} + + \noindent Isar type classes also directly support code generation + in a Haskell like fashion. + + This tutorial demonstrates common elements of structured specifications + and abstract reasoning with type classes by the algebraic hierarchy of + semigroups, monoids and groups. Our background theory is that of + Isabelle/HOL \cite{isa-tutorial}, for which some + familiarity is assumed. + + Here we merely present the look-and-feel for end users. + Internally, those are mapped to more primitive Isabelle concepts. + See \cite{Haftmann-Wenzel:2006:classes} for more detail.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{A simple algebra example \label{sec:example}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Class definition% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operator \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} that is + assumed to be associative:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ semigroup\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two + parts: the \qn{operational} part names the class parameter + (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them + (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and + \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, + yielding the global + parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the + global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Class instantiation \label{sec:class_inst}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The concrete type \isa{int} is made a \isa{semigroup} + instance by providing a suitable definition for the class parameter + \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}. + This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymotimes}\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} allows to define class parameters + at a particular instance using common specification tools (here, + \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} + opens a proof that the given parameters actually conform + to the class specification. Note that the first proof step + is the \hyperlink{method.default}{\mbox{\isa{default}}} method, + which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method. + This boils down an instance judgement to the relevant primitive + proof goals and should conveniently always be the first method applied + in an instantiation proof. + + From now on, the type-checker will consider \isa{int} + as a \isa{semigroup} automatically, i.e.\ any general results + are immediately available on concrete instances. + + \medskip Another instance of \isa{semigroup} are the natural numbers:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}m\ {\isasymotimes}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ m{\isacharparenright}\ auto\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} + in the primrec declaration; by default, the local name of + a class operation \isa{f} to instantiate on type constructor + \isa{{\isasymkappa}} are mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty, + these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command + or the corresponding ProofGeneral button.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Lifting and parametric types% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Overloaded definitions giving on class instantiation + may include recursion over the syntactic structure of types. + As a canonical example, we model product semigroups + using our simple algebra:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ fst\ p\isactrlisub {\isadigit{2}}{\isacharcomma}\ snd\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ snd\ p\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ p\isactrlisub {\isadigit{1}}\ p\isactrlisub {\isadigit{2}}\ p\isactrlisub {\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}\ {\isacharequal}\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\ \ \ \ \ \ \isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Associativity from product semigroups is + established using + the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical + associativity of the type components; these hypotheses + are facts due to the \isa{semigroup} constraints imposed + on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition. + Indeed, this pattern often occurs with parametric types + and type classes.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Subclassing% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral) + by extending \isa{semigroup} + with one additional parameter \isa{neutral} together + with its property:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Again, we prove some instances, by + providing suitable parameter definitions and proofs for the + additional specifications. Observe that instantiations + for types with the same arity may be simultaneous:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ neutral{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ neutral{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isacharcomma}\ {\isasymone}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoidl\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoidl{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutl{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Fully-fledged monoids are modelled by another subclass + which does not add new parameters but tightens the specification:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline +\ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoid\ \isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\ \isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoid\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoid{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}p\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutr{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent To finish our small algebra example, we add a \isa{group} class + with a corresponding instance:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ inverse{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isacharequal}\ {\isacharminus}\ {\isacharparenleft}i{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}int{\isacharunderscore}def\ neutral{\isacharunderscore}int{\isacharunderscore}def\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsection{Type classes as locales% +} +\isamarkuptrue% +% +\isamarkupsubsection{A look behind the scene% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The example above gives an impression how Isar type classes work + in practice. As stated in the introduction, classes also provide + a link to Isar's locale system. Indeed, the logical core of a class + is nothing else than a locale:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ idem\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent essentially introduces the locale% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +\ % +\endisadeliminvisible +% +\isataginvisible +\isacommand{setup}\isamarkupfalse% +\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +\isanewline +% +\isadelimquote +\isanewline +% +\endisadelimquote +% +\isatagquote +\isacommand{locale}\isamarkupfalse% +\ idem\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent together with corresponding constant(s):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{consts}\isamarkupfalse% +\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The connection to the type system is done by means + of a primitive axclass% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +\ % +\endisadeliminvisible +% +\isataginvisible +\isacommand{setup}\isamarkupfalse% +\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +\isanewline +% +\isadelimquote +\isanewline +% +\endisadelimquote +% +\isatagquote +\isacommand{axclass}\isamarkupfalse% +\ idem\ {\isacharless}\ type\isanewline +\ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isadeliminvisible +\ % +\endisadeliminvisible +% +\isataginvisible +\isacommand{setup}\isamarkupfalse% +\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\begin{isamarkuptext}% +\noindent together with a corresponding interpretation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{interpretation}\isamarkupfalse% +\ idem{\isacharunderscore}class{\isacharcolon}\isanewline +\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ \isacommand{qed}\isamarkupfalse% +\ {\isacharparenleft}rule\ idem{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This gives you at hand the full power of the Isabelle module system; + conclusions in locale \isa{idem} are implicitly propagated + to class \isa{idem}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +\ % +\endisadeliminvisible +% +\isataginvisible +\isacommand{setup}\isamarkupfalse% +\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isamarkupsubsection{Abstract reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle locales enable reasoning at a general level, while results + are implicitly transferred to all instances. For example, we can + now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which + states that the function \isa{{\isacharparenleft}x\ {\isasymotimes}{\isacharparenright}} is injective:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ assoc\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target specification + indicates that the result is recorded within that context for later + use. This local theorem is also lifted to the global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of + \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Derived definitions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle locales support a concept of local definitions + in locales:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent If the locale \isa{group} is also a class, this local + definition is propagated onto a global definition of + \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} + with corresponding theorems + + \isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline% +pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}. + + \noindent As you can see from this example, for local + definitions you may use any specification tool + which works together with locales (e.g. \cite{krauss2006}).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{A functor analogy% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We introduced Isar classes by analogy to type classes + functional programming; if we reconsider this in the + context of what has been said about type classes and locales, + we can drive this analogy further by stating that type + classes essentially correspond to functors which have + a canonical interpretation as type classes. + Anyway, there is also the possibility of other interpretations. + For example, also \isa{list}s form a monoid with + \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it + seems inappropriate to apply to lists + the same operations as for genuinely algebraic types. + In such a case, we simply can do a particular interpretation + of monoids for lists:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{interpretation}\isamarkupfalse% +\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \isacommand{qed}\isamarkupfalse% +\ auto% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This enables us to apply facts on monoids + to lists, e.g. \isa{{\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ x\ {\isacharequal}\ x}. + + When using this interpretation pattern, it may also + be appropriate to map derived definitions accordingly:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ replicate\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ list\ {\isasymRightarrow}\ {\isasymalpha}\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{interpretation}\isamarkupfalse% +\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{interpret}\isamarkupfalse% +\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ n\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ n\ {\isacharequal}\ replicate\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ n{\isacharparenright}\ auto\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\ intro{\isacharunderscore}locales% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsubsection{Additional subclass relations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Any \isa{group} is also a \isa{monoid}; this + can be made explicit by claiming an additional + subclass relation, + together with a proof of the logical difference:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{subclass}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ monoid\isanewline +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ invl\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}\ neutl\ invl\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ {\isasymone}{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ left{\isacharunderscore}cancel\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The logical proof is carried out on the locale level. + Afterwards it is propagated + to the type system, making \isa{group} an instance of + \isa{monoid} by adding an additional edge + to the graph of subclass relations + (cf.\ \figref{fig:subclass}). + + \begin{figure}[htbp] + \begin{center} + \small + \unitlength 0.6mm + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){\isa{semigroup}}} + \put(20,40){\makebox(0,0){\isa{monoidl}}} + \put(00,20){\makebox(0,0){\isa{monoid}}} + \put(40,00){\makebox(0,0){\isa{group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(25,35){\vector(1,-3){10}} + \end{picture} + \hspace{8em} + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){\isa{semigroup}}} + \put(20,40){\makebox(0,0){\isa{monoidl}}} + \put(00,20){\makebox(0,0){\isa{monoid}}} + \put(40,00){\makebox(0,0){\isa{group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(05,15){\vector(3,-1){30}} + \end{picture} + \caption{Subclass relationship of monoids and groups: + before and after establishing the relationship + \isa{group\ {\isasymsubseteq}\ monoid}; transitive edges left out.} + \label{fig:subclass} + \end{center} + \end{figure} +7 + For illustration, a derived definition + in \isa{group} which uses \isa{pow{\isacharunderscore}nat}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline +\ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline +\ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent yields the global definition of + \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} + with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{A note on syntax% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As a commodity, class context syntax allows to refer + to local class operations and their global counterparts + uniformly; type inference resolves ambiguities. For example:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{context}\isamarkupfalse% +\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{term}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % +\isamarkupcmt{example 1% +} +\isanewline +\isacommand{term}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % +\isamarkupcmt{example 2% +} +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{term}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % +\isamarkupcmt{example 3% +} +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Here in example 1, the term refers to the local class operation + \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint + enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}. + In the global context in example 3, the reference is + to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Further issues% +} +\isamarkuptrue% +% +\isamarkupsubsection{Type classes and code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Turning back to the first motivation for type classes, + namely overloading, it is obvious that overloading + stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and + \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} + targets naturally maps to Haskell type classes. + The code generator framework \cite{isabelle-codegen} + takes this into account. Concerning target languages + lacking type classes (e.g.~SML), type classes + are implemented by explicit dictionary construction. + As example, let's go back to the power function:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ example\ {\isacharcolon}{\isacharcolon}\ int\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This maps to Haskell as:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}module Example where {\char123}\\ +\hspace*{0pt}\\ +\hspace*{0pt}\\ +\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}nat{\char95}aux ::~Integer -> Nat -> Nat;\\ +\hspace*{0pt}nat{\char95}aux i n = (if i <= 0 then n else nat{\char95}aux (i - 1) (Suc n));\\ +\hspace*{0pt}\\ +\hspace*{0pt}nat ::~Integer -> Nat;\\ +\hspace*{0pt}nat i = nat{\char95}aux i Zero{\char95}nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}class Semigroup a where {\char123}\\ +\hspace*{0pt} ~mult ::~a -> a -> a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Semigroup a) => Monoidl a where {\char123}\\ +\hspace*{0pt} ~neutral ::~a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Monoidl a) => Monoid a where {\char123}\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Monoid a) => Group a where {\char123}\\ +\hspace*{0pt} ~inverse ::~a -> a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}inverse{\char95}int ::~Integer -> Integer;\\ +\hspace*{0pt}inverse{\char95}int i = negate i;\\ +\hspace*{0pt}\\ +\hspace*{0pt}neutral{\char95}int ::~Integer;\\ +\hspace*{0pt}neutral{\char95}int = 0;\\ +\hspace*{0pt}\\ +\hspace*{0pt}mult{\char95}int ::~Integer -> Integer -> Integer;\\ +\hspace*{0pt}mult{\char95}int i j = i + j;\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Semigroup Integer where {\char123}\\ +\hspace*{0pt} ~mult = mult{\char95}int;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Monoidl Integer where {\char123}\\ +\hspace*{0pt} ~neutral = neutral{\char95}int;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Monoid Integer where {\char123}\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Group Integer where {\char123}\\ +\hspace*{0pt} ~inverse = inverse{\char95}int;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\ +\hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\ +\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\ +\hspace*{0pt}\\ +\hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\ +\hspace*{0pt}pow{\char95}int k x =\\ +\hspace*{0pt} ~(if 0 <= k then pow{\char95}nat (nat k) x\\ +\hspace*{0pt} ~~~else inverse (pow{\char95}nat (nat (negate k)) x));\\ +\hspace*{0pt}\\ +\hspace*{0pt}example ::~Integer;\\ +\hspace*{0pt}example = pow{\char95}int 10 (-2);\\ +\hspace*{0pt}\\ +\hspace*{0pt}{\char125}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The whole code in SML with explicit dictionary passing:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun nat{\char95}aux i n =\\ +\hspace*{0pt} ~(if IntInf.<= (i,~(0 :~IntInf.int)) then n\\ +\hspace*{0pt} ~~~else nat{\char95}aux (IntInf.- (i,~(1 :~IntInf.int))) (Suc n));\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun nat i = nat{\char95}aux i Zero{\char95}nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ +\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a monoidl =\\ +\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl :~'a semigroup,~neutral :~'a{\char125};\\ +\hspace*{0pt}fun semigroup{\char95}monoidl (A{\char95}:'a monoidl) = {\char35}Classes{\char95}{\char95}semigroup{\char95}monoidl A{\char95};\\ +\hspace*{0pt}fun neutral (A{\char95}:'a monoidl) = {\char35}neutral A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a monoid = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid :~'a monoidl{\char125};\\ +\hspace*{0pt}fun monoidl{\char95}monoid (A{\char95}:'a monoid) = {\char35}Classes{\char95}{\char95}monoidl{\char95}monoid A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a group = {\char123}Classes{\char95}{\char95}monoid{\char95}group :~'a monoid,~inverse :~'a -> 'a{\char125};\\ +\hspace*{0pt}fun monoid{\char95}group (A{\char95}:'a group) = {\char35}Classes{\char95}{\char95}monoid{\char95}group A{\char95};\\ +\hspace*{0pt}fun inverse (A{\char95}:'a group) = {\char35}inverse A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ +\hspace*{0pt}\\ +\hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoidl{\char95}int =\\ +\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int monoidl;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoid{\char95}int = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int monoid;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val group{\char95}int =\\ +\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int group;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\ +\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\ +\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun pow{\char95}int A{\char95}~k x =\\ +\hspace*{0pt} ~(if IntInf.<= ((0 :~IntInf.int),~k)\\ +\hspace*{0pt} ~~~then pow{\char95}nat (monoid{\char95}group A{\char95}) (nat k) x\\ +\hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\ +\hspace*{0pt}\\ +\hspace*{0pt}val example :~IntInf.int =\\ +\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsubsection{Inspecting the type class universe% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +To facilitate orientation in complex subclass structures, + two diagnostics commands are provided: + + \begin{description} + + \item[\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isacharunderscore}classes}}}}] print a list of all classes + together with associated operations etc. + + \item[\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isacharunderscore}deps}}}}] visualizes the subclass relation + between all classes as a Hasse diagram. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/classes.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/classes.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,50 @@ + +\documentclass[12pt,a4paper,fleqn]{report} +\usepackage{latexsym,graphicx} +\usepackage[refpage]{nomencl} +\usepackage{../iman,../extra,../isar,../proof} +\usepackage{../isabelle,../isabellesym} +\usepackage{style} +\usepackage{../pdfsetup} + + +\hyphenation{Isabelle} +\hyphenation{Isar} +\isadroptag{theory} + +\title{\includegraphics[scale=0.5]{isabelle_isar} + \\[4ex] Haskell-style type classes with Isabelle/Isar} +\author{\emph{Florian Haftmann}} + +\begin{document} + +\maketitle + +\begin{abstract} + This tutorial introduces the look-and-feel of Isar type classes + to the end-user; Isar type classes are a convenient mechanism + for organizing specifications, overcoming some drawbacks + of raw axiomatic type classes. Essentially, they combine + an operational aspect (in the manner of Haskell) with + a logical aspect, both managed uniformly. +\end{abstract} + +\thispagestyle{empty}\clearpage + +\pagenumbering{roman} +\clearfirst + +\input{Thy/document/Classes.tex} + +\begingroup +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{../manual} +\endgroup + +\end{document} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Classes/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/style.sty Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,43 @@ + +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} +\newcommand{\figref}[1]{figure~\ref{#1}} + +%% logical markup +\newcommand{\strong}[1]{{\bfseries {#1}}} +\newcommand{\qn}[1]{\emph{#1}} + +%% typographic conventions +\newcommand{\qt}[1]{``{#1}''} + +%% verbatim text +\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} + +%% quoted segments +\makeatletter +\isakeeptag{quote} +\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}} +\renewcommand{\isatagquote}{\begin{quotesegment}} +\renewcommand{\endisatagquote}{\end{quotesegment}} +\makeatother + +%% presentation +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +\pagestyle{headings} +\binperiod +\underscoreoff + +\renewcommand{\isadigit}[1]{\isamath{#1}} + +\isabellestyle{it} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "implementation" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/IsaMakefile Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,33 @@ + +## targets + +default: Thy +images: +test: Thy + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document + + +## Thy + +THY = $(LOG)/HOL-Thy.gz + +Thy: $(THY) + +$(THY): Thy/ROOT.ML Thy/*.thy ../antiquote_setup.ML ../more_antiquote.ML + @$(USEDIR) HOL Thy + + +## clean + +clean: + @rm -f $(THY) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Makefile Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,35 @@ + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = codegen + +FILES = $(NAME).tex Thy/document/*.tex \ + style.sty ../iman.sty ../extra.sty ../isar.sty \ + ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ + ../manual.bib ../proof.sty + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_isar.eps codegen_process.ps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_isar.pdf codegen_process.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/Adaption.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Adaption.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,361 @@ +theory Adaption +imports Setup +begin + +setup %invisible {* Code_Target.extend_target ("\", ("SML", K I)) *} + +section {* Adaption to target languages \label{sec:adaption} *} + +subsection {* Adapting code generation *} + +text {* + The aspects of code generation introduced so far have two aspects + in common: + + \begin{itemize} + \item They act uniformly, without reference to a specific + target language. + \item They are \emph{safe} in the sense that as long as you trust + the code generator meta theory and implementation, you cannot + produce programs that yield results which are not derivable + in the logic. + \end{itemize} + + \noindent In this section we will introduce means to \emph{adapt} the serialiser + to a specific target language, i.e.~to print program fragments + in a way which accommodates \qt{already existing} ingredients of + a target language environment, for three reasons: + + \begin{itemize} + \item improving readability and aesthetics of generated code + \item gaining efficiency + \item interface with language parts which have no direct counterpart + in @{text "HOL"} (say, imperative data structures) + \end{itemize} + + \noindent Generally, you should avoid using those features yourself + \emph{at any cost}: + + \begin{itemize} + \item The safe configuration methods act uniformly on every target language, + whereas for adaption you have to treat each target language separate. + \item Application is extremely tedious since there is no abstraction + which would allow for a static check, making it easy to produce garbage. + \item More or less subtle errors can be introduced unconsciously. + \end{itemize} + + \noindent However, even if you ought refrain from setting up adaption + yourself, already the @{text "HOL"} comes with some reasonable default + adaptions (say, using target language list syntax). There also some + common adaption cases which you can setup by importing particular + library theories. In order to understand these, we provide some clues here; + these however are not supposed to replace a careful study of the sources. +*} + +subsection {* The adaption principle *} + +text {* + The following figure illustrates what \qt{adaption} is conceptually + supposed to be: + + \begin{figure}[here] + \begin{tikzpicture}[scale = 0.5] + \tikzstyle water=[color = blue, thick] + \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] + \tikzstyle process=[color = green, semithick, ->] + \tikzstyle adaption=[color = red, semithick, ->] + \tikzstyle target=[color = black] + \foreach \x in {0, ..., 24} + \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin + + (0.25, -0.25) cos + (0.25, 0.25); + \draw[style=ice] (1, 0) -- + (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; + \draw[style=ice] (9, 0) -- + (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; + \draw[style=ice] (15, -6) -- + (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; + \draw[style=process] + (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); + \draw[style=process] + (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); + \node (adaption) at (11, -2) [style=adaption] {adaption}; + \node at (19, 3) [rotate=90] {generated}; + \node at (19.5, -5) {language}; + \node at (19.5, -3) {library}; + \node (includes) at (19.5, -1) {includes}; + \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 + \draw[style=process] + (includes) -- (serialisation); + \draw[style=process] + (reserved) -- (serialisation); + \draw[style=adaption] + (adaption) -- (serialisation); + \draw[style=adaption] + (adaption) -- (includes); + \draw[style=adaption] + (adaption) -- (reserved); + \end{tikzpicture} + \caption{The adaption principle} + \label{fig:adaption} + \end{figure} + + \noindent In the tame view, code generation acts as broker between + @{text logic}, @{text "intermediate language"} and + @{text "target language"} by means of @{text translation} and + @{text serialisation}; for the latter, the serialiser has to observe + the structure of the @{text language} itself plus some @{text reserved} + keywords which have to be avoided for generated code. + However, if you consider @{text adaption} mechanisms, the code generated + by the serializer is just the tip of the iceberg: + + \begin{itemize} + \item @{text serialisation} can be \emph{parametrised} such that + logical entities are mapped to target-specific ones + (e.g. target-specific list syntax, + see also \secref{sec:adaption_mechanisms}) + \item Such parametrisations can involve references to a + target-specific standard @{text library} (e.g. using + the @{text Haskell} @{verbatim Maybe} type instead + of the @{text HOL} @{type "option"} type); + if such are used, the corresponding identifiers + (in our example, @{verbatim Maybe}, @{verbatim Nothing} + and @{verbatim Just}) also have to be considered @{text reserved}. + \item Even more, the user can enrich the library of the + target-language by providing code snippets + (\qt{@{text "includes"}}) which are prepended to + any generated code (see \secref{sec:include}); this typically + also involves further @{text reserved} identifiers. + \end{itemize} + + \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms + have to act consistently; it is at the discretion of the user + to take care for this. +*} + +subsection {* Common adaption patterns *} + +text {* + The @{theory HOL} @{theory Main} theory already provides a code + generator setup + which should be suitable for most applications. Common extensions + and modifications are available by certain theories of the @{text HOL} + library; beside being useful in applications, they may serve + as a tutorial for customising the code generator setup (see below + \secref{sec:adaption_mechanisms}). + + \begin{description} + + \item[@{theory "Code_Integer"}] represents @{text HOL} integers by big + integer literals in target languages. + \item[@{theory "Code_Char"}] represents @{text HOL} characters by + character literals in target languages. + \item[@{theory "Code_Char_chr"}] like @{text "Code_Char"}, + but also offers treatment of character codes; includes + @{theory "Code_Char"}. + \item[@{theory "Efficient_Nat"}] \label{eff_nat} implements natural numbers by integers, + which in general will result in higher efficiency; pattern + matching with @{term "0\nat"} / @{const "Suc"} + is eliminated; includes @{theory "Code_Integer"} + and @{theory "Code_Index"}. + \item[@{theory "Code_Index"}] provides an additional datatype + @{typ index} which is mapped to target-language built-in integers. + Useful for code setups which involve e.g. indexing of + target-language arrays. + \item[@{theory "Code_Message"}] provides an additional datatype + @{typ message_string} which is isomorphic to strings; + @{typ message_string}s are mapped to target-language strings. + Useful for code setups which involve e.g. printing (error) messages. + + \end{description} + + \begin{warn} + When importing any of these theories, they should form the last + items in an import list. Since these theories adapt the + code generator setup in a non-conservative fashion, + strange effects may occur otherwise. + \end{warn} +*} + + +subsection {* Parametrising serialisation \label{sec:adaption_mechanisms} *} + +text {* + Consider the following function and its corresponding + SML code: +*} + +primrec %quote in_interval :: "nat \ nat \ nat \ bool" where + "in_interval (k, l) n \ k \ n \ n \ l" +(*<*) +code_type %invisible bool + (SML) +code_const %invisible True and False and "op \" and Not + (SML and and and) +(*>*) +text %quote {*@{code_stmts in_interval (SML)}*} + +text {* + \noindent Though this is correct code, it is a little bit unsatisfactory: + boolean values and operators are materialised as distinguished + entities with have nothing to do with the SML-built-in notion + of \qt{bool}. This results in less readable code; + additionally, eager evaluation may cause programs to + loop or break which would perfectly terminate when + the existing SML @{verbatim "bool"} would be used. To map + the HOL @{typ bool} on SML @{verbatim "bool"}, we may use + \qn{custom serialisations}: +*} + +code_type %quotett bool + (SML "bool") +code_const %quotett True and False and "op \" + (SML "true" and "false" and "_ andalso _") + +text {* + \noindent The @{command code_type} command takes a type constructor + as arguments together with a list of custom serialisations. + Each custom serialisation starts with a target language + identifier followed by an expression, which during + code serialisation is inserted whenever the type constructor + would occur. For constants, @{command code_const} implements + the corresponding mechanism. Each ``@{verbatim "_"}'' in + a serialisation expression is treated as a placeholder + for the type constructor's (the constant's) arguments. +*} + +text %quote {*@{code_stmts in_interval (SML)}*} + +text {* + \noindent This still is not perfect: the parentheses + around the \qt{andalso} expression are superfluous. + Though the serialiser + by no means attempts to imitate the rich Isabelle syntax + framework, it provides some common idioms, notably + associative infixes with precedences which may be used here: +*} + +code_const %quotett "op \" + (SML infixl 1 "andalso") + +text %quote {*@{code_stmts in_interval (SML)}*} + +text {* + \noindent The attentive reader may ask how we assert that no generated + code will accidentally overwrite. For this reason the serialiser has + an internal table of identifiers which have to be avoided to be used + for new declarations. Initially, this table typically contains the + keywords of the target language. It can be extended manually, thus avoiding + accidental overwrites, using the @{command "code_reserved"} command: +*} + +code_reserved %quote "\" bool true false andalso + +text {* + \noindent Next, we try to map HOL pairs to SML pairs, using the + infix ``@{verbatim "*"}'' type constructor and parentheses: +*} +(*<*) +code_type %invisible * + (SML) +code_const %invisible Pair + (SML) +(*>*) +code_type %quotett * + (SML infix 2 "*") +code_const %quotett Pair + (SML "!((_),/ (_))") + +text {* + \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser + never to put + parentheses around the whole expression (they are already present), + while the parentheses around argument place holders + tell not to put parentheses around the arguments. + The slash ``@{verbatim "/"}'' (followed by arbitrary white space) + inserts a space which may be used as a break if necessary + during pretty printing. + + These examples give a glimpse what mechanisms + custom serialisations provide; however their usage + requires careful thinking in order not to introduce + inconsistencies -- or, in other words: + custom serialisations are completely axiomatic. + + A further noteworthy details is that any special + character in a custom serialisation may be quoted + using ``@{verbatim "'"}''; thus, in + ``@{verbatim "fn '_ => _"}'' the first + ``@{verbatim "_"}'' is a proper underscore while the + second ``@{verbatim "_"}'' is a placeholder. +*} + + +subsection {* @{text Haskell} serialisation *} + +text {* + For convenience, the default + @{text HOL} setup for @{text Haskell} maps the @{class eq} class to + its counterpart in @{text Haskell}, giving custom serialisations + for the class @{class eq} (by command @{command code_class}) and its operation + @{const HOL.eq} +*} + +code_class %quotett eq + (Haskell "Eq") + +code_const %quotett "op =" + (Haskell infixl 4 "==") + +text {* + \noindent A problem now occurs whenever a type which + is an instance of @{class eq} in @{text HOL} is mapped + on a @{text Haskell}-built-in type which is also an instance + of @{text Haskell} @{text Eq}: +*} + +typedecl %quote bar + +instantiation %quote bar :: eq +begin + +definition %quote "eq_class.eq (x\bar) y \ x = y" + +instance %quote by default (simp add: eq_bar_def) + +end %quote + +code_type %quotett bar + (Haskell "Integer") + +text {* + \noindent The code generator would produce + an additional instance, which of course is rejected by the @{text Haskell} + compiler. + To suppress this additional instance, use + @{text "code_instance"}: +*} + +code_instance %quotett bar :: eq + (Haskell -) + + +subsection {* Enhancing the target language context \label{sec:include} *} + +text {* + In rare cases it is necessary to \emph{enrich} the context of a + target language; this is accomplished using the @{command "code_include"} + command: +*} + +code_include %quotett Haskell "Errno" +{*errno i = error ("Error number: " ++ show i)*} + +code_reserved %quotett Haskell Errno + +text {* + \noindent Such named @{text include}s are then prepended to every generated code. + Inspect such code in order to find out how @{command "code_include"} behaves + with respect to a particular target language. +*} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/Codegen.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Codegen.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,2 @@ + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/Further.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Further.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,113 @@ +theory Further +imports Setup +begin + +section {* Further issues \label{sec:further} *} + +subsection {* Further reading *} + +text {* + Do dive deeper into the issue of code generation, you should visit + the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which + contains exhaustive syntax diagrams. +*} + +subsection {* Modules *} + +text {* + When invoking the @{command export_code} command it is possible to leave + out the @{keyword "module_name"} part; then code is distributed over + different modules, where the module name space roughly is induced + by the @{text Isabelle} theory name space. + + Then sometimes the awkward situation occurs that dependencies between + definitions introduce cyclic dependencies between modules, which in the + @{text Haskell} world leaves you to the mercy of the @{text Haskell} implementation + you are using, while for @{text SML}/@{text OCaml} code generation is not possible. + + A solution is to declare module names explicitly. + Let use assume the three cyclically dependent + modules are named \emph{A}, \emph{B} and \emph{C}. + Then, by stating +*} + +code_modulename %quote SML + A ABC + B ABC + C ABC + +text {* + we explicitly map all those modules on \emph{ABC}, + resulting in an ad-hoc merge of this three modules + at serialisation time. +*} + +subsection {* Evaluation oracle *} + +text {* + Code generation may also be used to \emph{evaluate} expressions + (using @{text SML} as target language of course). + For instance, the @{command value} allows to reduce an expression to a + normal form with respect to the underlying code equations: +*} + +value %quote "42 / (12 :: rat)" + +text {* + \noindent will display @{term "7 / (2 :: rat)"}. + + The @{method eval} method tries to reduce a goal by code generation to @{term True} + and solves it in that case, but fails otherwise: +*} + +lemma %quote "42 / (12 :: rat) = 7 / 2" + by %quote eval + +text {* + \noindent The soundness of the @{method eval} method depends crucially + on the correctness of the code generator; this is one of the reasons + why you should not use adaption (see \secref{sec:adaption}) frivolously. +*} + +subsection {* Code antiquotation *} + +text {* + In scenarios involving techniques like reflection it is quite common + that code generated from a theory forms the basis for implementing + a proof procedure in @{text SML}. To facilitate interfacing of generated code + with system code, the code generator provides a @{text code} antiquotation: +*} + +datatype %quote form = T | F | And form form | Or form form + +ML %quote {* + fun eval_form @{code T} = true + | eval_form @{code F} = false + | eval_form (@{code And} (p, q)) = + eval_form p andalso eval_form q + | eval_form (@{code Or} (p, q)) = + eval_form p orelse eval_form q; +*} + +text {* + \noindent @{text code} takes as argument the name of a constant; after the + whole @{text SML} is read, the necessary code is generated transparently + and the corresponding constant names are inserted. This technique also + allows to use pattern matching on constructors stemming from compiled + @{text datatypes}. + + For a less simplistic example, theory @{theory Ferrack} is + a good reference. +*} + +subsection {* Imperative data structures *} + +text {* + If you consider imperative data structures as inevitable for a specific + application, you should consider + \emph{Imperative Functional Programming with Isabelle/HOL} + (\cite{bulwahn-et-al:2008:imperative}); + the framework described there is available in theory @{theory Imperative_HOL}. +*} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/Introduction.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Introduction.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,206 @@ +theory Introduction +imports Setup +begin + +chapter {* Code generation from @{text "Isabelle/HOL"} theories *} + +section {* Introduction and Overview *} + +text {* + This tutorial introduces a generic code generator for the + @{text Isabelle} system. + Generic in the sense that the + \qn{target language} for which code shall ultimately be + generated is not fixed but may be an arbitrary state-of-the-art + functional programming language (currently, the implementation + supports @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and @{text Haskell} + \cite{haskell-revised-report}). + + Conceptually the code generator framework is part + of Isabelle's @{theory Pure} meta logic framework; the logic + @{theory HOL} which is an extension of @{theory Pure} + already comes with a reasonable framework setup and thus provides + a good working horse for raising code-generation-driven + applications. So, we assume some familiarity and experience + with the ingredients of the @{theory HOL} distribution theories. + (see also \cite{isa-tutorial}). + + The code generator aims to be usable with no further ado + in most cases while allowing for detailed customisation. + This manifests in the structure of this tutorial: after a short + conceptual introduction with an example (\secref{sec:intro}), + we discuss the generic customisation facilities (\secref{sec:program}). + A further section (\secref{sec:adaption}) is dedicated to the matter of + \qn{adaption} to specific target language environments. After some + further issues (\secref{sec:further}) we conclude with an overview + of some ML programming interfaces (\secref{sec:ml}). + + \begin{warn} + Ultimately, the code generator which this tutorial deals with + is supposed to replace the existing code generator + by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. + So, for the moment, there are two distinct code generators + in Isabelle. In case of ambiguity, we will refer to the framework + described here as @{text "generic code generator"}, to the + other as @{text "SML code generator"}. + Also note that while the framework itself is + object-logic independent, only @{theory HOL} provides a reasonable + framework setup. + \end{warn} + +*} + +subsection {* Code generation via shallow embedding \label{sec:intro} *} + +text {* + The key concept for understanding @{text Isabelle}'s code generation is + \emph{shallow embedding}, i.e.~logical entities like constants, types and + classes are identified with corresponding concepts in the target language. + + Inside @{theory HOL}, the @{command datatype} and + @{command definition}/@{command primrec}/@{command fun} declarations form + the core of a functional programming language. The default code generator setup + allows to turn those into functional programs immediately. + This means that \qt{naive} code generation can proceed without further ado. + For example, here a simple \qt{implementation} of amortised queues: +*} + +datatype %quote 'a queue = AQueue "'a list" "'a list" + +definition %quote empty :: "'a queue" where + "empty = AQueue [] []" + +primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where + "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" + +fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where + "dequeue (AQueue [] []) = (None, AQueue [] [])" + | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" + | "dequeue (AQueue xs []) = + (case rev xs of y # ys \ (Some y, AQueue [] ys))" + +text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *} + +export_code %quote empty dequeue enqueue in SML + module_name Example file "examples/example.ML" + +text {* \noindent resulting in the following code: *} + +text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} + +text {* + \noindent The @{command export_code} command takes a space-separated list of + constants for which code shall be generated; anything else needed for those + is added implicitly. Then follows a target language identifier + (@{text SML}, @{text OCaml} or @{text Haskell}) and a freely chosen module name. + A file name denotes the destination to store the generated code. Note that + the semantics of the destination depends on the target language: for + @{text SML} and @{text OCaml} it denotes a \emph{file}, for @{text Haskell} + it denotes a \emph{directory} where a file named as the module name + (with extension @{text ".hs"}) is written: +*} + +export_code %quote empty dequeue enqueue in Haskell + module_name Example file "examples/" + +text {* + \noindent This is how the corresponding code in @{text Haskell} looks like: +*} + +text %quote {*@{code_stmts empty enqueue dequeue (Haskell)}*} + +text {* + \noindent This demonstrates the basic usage of the @{command export_code} command; + for more details see \secref{sec:further}. +*} + +subsection {* Code generator architecture \label{sec:concept} *} + +text {* + What you have seen so far should be already enough in a lot of cases. If you + are content with this, you can quit reading here. Anyway, in order to customise + and adapt the code generator, it is inevitable to gain some understanding + how it works. + + \begin{figure}[h] + \begin{tikzpicture}[x = 4.2cm, y = 1cm] + \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; + \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; + \tikzstyle process_arrow=[->, semithick, color = green]; + \node (HOL) at (0, 4) [style=entity] {@{text "Isabelle/HOL"} theory}; + \node (eqn) at (2, 2) [style=entity] {code equations}; + \node (iml) at (2, 0) [style=entity] {intermediate language}; + \node (seri) at (1, 0) [style=process] {serialisation}; + \node (SML) at (0, 3) [style=entity] {@{text SML}}; + \node (OCaml) at (0, 2) [style=entity] {@{text OCaml}}; + \node (further) at (0, 1) [style=entity] {@{text "\"}}; + \node (Haskell) at (0, 0) [style=entity] {@{text Haskell}}; + \draw [style=process_arrow] (HOL) .. controls (2, 4) .. + node [style=process, near start] {selection} + node [style=process, near end] {preprocessing} + (eqn); + \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); + \draw [style=process_arrow] (iml) -- (seri); + \draw [style=process_arrow] (seri) -- (SML); + \draw [style=process_arrow] (seri) -- (OCaml); + \draw [style=process_arrow, dashed] (seri) -- (further); + \draw [style=process_arrow] (seri) -- (Haskell); + \end{tikzpicture} + \caption{Code generator architecture} + \label{fig:arch} + \end{figure} + + The code generator employs a notion of executability + for three foundational executable ingredients known + from functional programming: + \emph{code equations}, \emph{datatypes}, and + \emph{type classes}. A code equation as a first approximation + is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \ t\<^isub>n \ t"} + (an equation headed by a constant @{text f} with arguments + @{text "t\<^isub>1 t\<^isub>2 \ t\<^isub>n"} and right hand side @{text t}). + Code generation aims to turn code equations + into a functional program. This is achieved by three major + components which operate sequentially, i.e. the result of one is + the input + of the next in the chain, see diagram \ref{fig:arch}: + + \begin{itemize} + + \item Out of the vast collection of theorems proven in a + \qn{theory}, a reasonable subset modelling + code equations is \qn{selected}. + + \item On those selected theorems, certain + transformations are carried out + (\qn{preprocessing}). Their purpose is to turn theorems + representing non- or badly executable + specifications into equivalent but executable counterparts. + The result is a structured collection of \qn{code theorems}. + + \item Before the selected code equations are continued with, + they can be \qn{preprocessed}, i.e. subjected to theorem + transformations. This \qn{preprocessor} is an interface which + allows to apply + the full expressiveness of ML-based theorem transformations + to code generation; motivating examples are shown below, see + \secref{sec:preproc}. + The result of the preprocessing step is a structured collection + of code equations. + + \item These code equations are \qn{translated} to a program + in an abstract intermediate language. Think of it as a kind + of \qt{Mini-Haskell} with four \qn{statements}: @{text data} + (for datatypes), @{text fun} (stemming from code equations), + also @{text class} and @{text inst} (for type classes). + + \item Finally, the abstract program is \qn{serialised} into concrete + source code of a target language. + + \end{itemize} + + \noindent From these steps, only the two last are carried out outside the logic; by + keeping this layer as thin as possible, the amount of code to trust is + kept to a minimum. +*} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/ML.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/ML.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,177 @@ +theory "ML" +imports Setup +begin + +section {* ML system interfaces \label{sec:ml} *} + +text {* + Since the code generator framework not only aims to provide + a nice Isar interface but also to form a base for + code-generation-based applications, here a short + description of the most important ML interfaces. +*} + +subsection {* Executable theory content: @{text Code} *} + +text {* + This Pure module implements the core notions of + executable content of a theory. +*} + +subsubsection {* Managing executable content *} + +text %mlref {* + \begin{mldecls} + @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ + @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ + @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\ + @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\ + @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\ + @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option) + -> theory -> theory"} \\ + @{index_ML Code.del_functrans: "string -> theory -> theory"} \\ + @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ + @{index_ML Code.get_datatype: "theory -> string + -> (string * sort) list * (string * typ list) list"} \\ + @{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"} + \end{mldecls} + + \begin{description} + + \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function + theorem @{text "thm"} to executable content. + + \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function + theorem @{text "thm"} from executable content, if present. + + \item @{ML Code.add_eqnl}~@{text "(const, lthms)"}~@{text "thy"} adds + suspended code equations @{text lthms} for constant + @{text const} to executable content. + + \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes + the preprocessor simpset. + + \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds + function transformer @{text f} (named @{text name}) to executable content; + @{text f} is a transformer of the code equations belonging + to a certain function definition, depending on the + current theory context. Returning @{text NONE} indicates that no + transformation took place; otherwise, the whole process will be iterated + with the new code equations. + + \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes + function transformer named @{text name} from executable content. + + \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds + a datatype to executable content, with generation + set @{text cs}. + + \item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"} + returns type constructor corresponding to + constructor @{text const}; returns @{text NONE} + if @{text const} is no constructor. + + \end{description} +*} + +subsection {* Auxiliary *} + +text %mlref {* + \begin{mldecls} + @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\ + @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\ + @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML Code_Unit.read_const}~@{text thy}~@{text s} + reads a constant as a concrete term expression @{text s}. + + \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm} + extracts the constant and its type from a code equation @{text thm}. + + \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm} + rewrites a code equation @{text thm} with a simpset @{text ss}; + only arguments and right hand side are rewritten, + not the head of the code equation. + + \end{description} + +*} + +subsection {* Implementing code generator applications *} + +text {* + Implementing code generator applications on top + of the framework set out so far usually not only + involves using those primitive interfaces + but also storing code-dependent data and various + other things. +*} + +subsubsection {* Data depending on the theory's executable content *} + +text {* + Due to incrementality of code generation, changes in the + theory's executable content have to be propagated in a + certain fashion. Additionally, such changes may occur + not only during theory extension but also during theory + merge, which is a little bit nasty from an implementation + point of view. The framework provides a solution + to this technical challenge by providing a functorial + data slot @{ML_functor CodeDataFun}; on instantiation + of this functor, the following types and operations + are required: + + \medskip + \begin{tabular}{l} + @{text "type T"} \\ + @{text "val empty: T"} \\ + @{text "val purge: theory \ string list option \ T \ T"} + \end{tabular} + + \begin{description} + + \item @{text T} the type of data to store. + + \item @{text empty} initial (empty) data. + + \item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content; + @{text consts} indicates the kind + of change: @{ML NONE} stands for a fundamental change + which invalidates any existing code, @{text "SOME consts"} + hints that executable content for constants @{text consts} + has changed. + + \end{description} + + \noindent An instance of @{ML_functor CodeDataFun} provides the following + interface: + + \medskip + \begin{tabular}{l} + @{text "get: theory \ T"} \\ + @{text "change: theory \ (T \ T) \ T"} \\ + @{text "change_yield: theory \ (T \ 'a * T) \ 'a * T"} + \end{tabular} + + \begin{description} + + \item @{text get} retrieval of the current data. + + \item @{text change} update of current data (cached!) + by giving a continuation. + + \item @{text change_yield} update with side result. + + \end{description} +*} + +text {* + \bigskip + + \emph{Happy proving, happy hacking!} +*} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/Program.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Program.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,534 @@ +theory Program +imports Introduction +begin + +section {* Turning Theories into Programs \label{sec:program} *} + +subsection {* The @{text "Isabelle/HOL"} default setup *} + +text {* + We have already seen how by default equations stemming from + @{command definition}/@{command primrec}/@{command fun} + statements are used for code generation. This default behaviour + can be changed, e.g. by providing different code equations. + All kinds of customisation shown in this section is \emph{safe} + in the sense that the user does not have to worry about + correctness -- all programs generatable that way are partially + correct. +*} + +subsection {* Selecting code equations *} + +text {* + Coming back to our introductory example, we + could provide an alternative code equations for @{const dequeue} + explicitly: +*} + +lemma %quote [code]: + "dequeue (AQueue xs []) = + (if xs = [] then (None, AQueue [] []) + else dequeue (AQueue [] (rev xs)))" + "dequeue (AQueue xs (y # ys)) = + (Some y, AQueue xs ys)" + by (cases xs, simp_all) (cases "rev xs", simp_all) + +text {* + \noindent The annotation @{text "[code]"} is an @{text Isar} + @{text attribute} which states that the given theorems should be + considered as code equations for a @{text fun} statement -- + the corresponding constant is determined syntactically. The resulting code: +*} + +text %quote {*@{code_stmts dequeue (consts) dequeue (Haskell)}*} + +text {* + \noindent You may note that the equality test @{term "xs = []"} has been + replaced by the predicate @{term "null xs"}. This is due to the default + setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). + + Changing the default constructor set of datatypes is also + possible. See \secref{sec:datatypes} for an example. + + As told in \secref{sec:concept}, code generation is based + on a structured collection of code theorems. + For explorative purpose, this collection + may be inspected using the @{command code_thms} command: +*} + +code_thms %quote dequeue + +text {* + \noindent prints a table with \emph{all} code equations + for @{const dequeue}, including + \emph{all} code equations those equations depend + on recursively. + + Similarly, the @{command code_deps} command shows a graph + visualising dependencies between code equations. +*} + +subsection {* @{text class} and @{text instantiation} *} + +text {* + Concerning type classes and code generation, let us examine an example + from abstract algebra: +*} + +class %quote semigroup = + fixes mult :: "'a \ 'a \ 'a" (infixl "\" 70) + assumes assoc: "(x \ y) \ z = x \ (y \ z)" + +class %quote monoid = semigroup + + fixes neutral :: 'a ("\") + assumes neutl: "\ \ x = x" + and neutr: "x \ \ = x" + +instantiation %quote nat :: monoid +begin + +primrec %quote mult_nat where + "0 \ n = (0\nat)" + | "Suc m \ n = n + m \ n" + +definition %quote neutral_nat where + "\ = Suc 0" + +lemma %quote add_mult_distrib: + fixes n m q :: nat + shows "(n + m) \ q = n \ q + m \ q" + by (induct n) simp_all + +instance %quote proof + fix m n q :: nat + show "m \ n \ q = m \ (n \ q)" + by (induct m) (simp_all add: add_mult_distrib) + show "\ \ n = n" + by (simp add: neutral_nat_def) + show "m \ \ = m" + by (induct m) (simp_all add: neutral_nat_def) +qed + +end %quote + +text {* + \noindent We define the natural operation of the natural numbers + on monoids: +*} + +primrec %quote (in monoid) pow :: "nat \ 'a \ 'a" where + "pow 0 a = \" + | "pow (Suc n) a = a \ pow n a" + +text {* + \noindent This we use to define the discrete exponentiation function: +*} + +definition %quote bexp :: "nat \ nat" where + "bexp n = pow n (Suc (Suc 0))" + +text {* + \noindent The corresponding code: +*} + +text %quote {*@{code_stmts bexp (Haskell)}*} + +text {* + \noindent This is a convenient place to show how explicit dictionary construction + manifests in generated code (here, the same example in @{text SML}): +*} + +text %quote {*@{code_stmts bexp (SML)}*} + +text {* + \noindent Note the parameters with trailing underscore (@{verbatim "A_"}) + which are the dictionary parameters. +*} + +subsection {* The preprocessor \label{sec:preproc} *} + +text {* + Before selected function theorems are turned into abstract + code, a chain of definitional transformation steps is carried + out: \emph{preprocessing}. In essence, the preprocessor + consists of two components: a \emph{simpset} and \emph{function transformers}. + + The \emph{simpset} allows to employ the full generality of the Isabelle + simplifier. Due to the interpretation of theorems + as code equations, rewrites are applied to the right + hand side and the arguments of the left hand side of an + equation, but never to the constant heading the left hand side. + An important special case are \emph{inline theorems} which may be + declared and undeclared using the + \emph{code inline} or \emph{code inline del} attribute respectively. + + Some common applications: +*} + +text_raw {* + \begin{itemize} +*} + +text {* + \item replacing non-executable constructs by executable ones: +*} + +lemma %quote [code inline]: + "x \ set xs \ x mem xs" by (induct xs) simp_all + +text {* + \item eliminating superfluous constants: +*} + +lemma %quote [code inline]: + "1 = Suc 0" by simp + +text {* + \item replacing executable but inconvenient constructs: +*} + +lemma %quote [code inline]: + "xs = [] \ List.null xs" by (induct xs) simp_all + +text_raw {* + \end{itemize} +*} + +text {* + \noindent \emph{Function transformers} provide a very general interface, + transforming a list of function theorems to another + list of function theorems, provided that neither the heading + constant nor its type change. The @{term "0\nat"} / @{const Suc} + pattern elimination implemented in + theory @{text Efficient_Nat} (see \secref{eff_nat}) uses this + interface. + + \noindent The current setup of the preprocessor may be inspected using + the @{command print_codesetup} command. + @{command code_thms} provides a convenient + mechanism to inspect the impact of a preprocessor setup + on code equations. + + \begin{warn} + The attribute \emph{code unfold} + associated with the @{text "SML code generator"} also applies to + the @{text "generic code generator"}: + \emph{code unfold} implies \emph{code inline}. + \end{warn} +*} + +subsection {* Datatypes \label{sec:datatypes} *} + +text {* + Conceptually, any datatype is spanned by a set of + \emph{constructors} of type @{text "\ = \ \ \ \\<^isub>1 \ \\<^isub>n"} where @{text + "{\\<^isub>1, \, \\<^isub>n}"} is exactly the set of \emph{all} type variables in + @{text "\"}. The HOL datatype package by default registers any new + datatype in the table of datatypes, which may be inspected using the + @{command print_codesetup} command. + + In some cases, it is appropriate to alter or extend this table. As + an example, we will develop an alternative representation of the + queue example given in \secref{sec:intro}. The amortised + representation is convenient for generating code but exposes its + \qt{implementation} details, which may be cumbersome when proving + theorems about it. Therefore, here a simple, straightforward + representation of queues: +*} + +datatype %quote 'a queue = Queue "'a list" + +definition %quote empty :: "'a queue" where + "empty = Queue []" + +primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where + "enqueue x (Queue xs) = Queue (xs @ [x])" + +fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where + "dequeue (Queue []) = (None, Queue [])" + | "dequeue (Queue (x # xs)) = (Some x, Queue xs)" + +text {* + \noindent This we can use directly for proving; for executing, + we provide an alternative characterisation: +*} + +definition %quote AQueue :: "'a list \ 'a list \ 'a queue" where + "AQueue xs ys = Queue (ys @ rev xs)" + +code_datatype %quote AQueue + +text {* + \noindent Here we define a \qt{constructor} @{const "AQueue"} which + is defined in terms of @{text "Queue"} and interprets its arguments + according to what the \emph{content} of an amortised queue is supposed + to be. Equipped with this, we are able to prove the following equations + for our primitive queue operations which \qt{implement} the simple + queues in an amortised fashion: +*} + +lemma %quote empty_AQueue [code]: + "empty = AQueue [] []" + unfolding AQueue_def empty_def by simp + +lemma %quote enqueue_AQueue [code]: + "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" + unfolding AQueue_def by simp + +lemma %quote dequeue_AQueue [code]: + "dequeue (AQueue xs []) = + (if xs = [] then (None, AQueue [] []) + else dequeue (AQueue [] (rev xs)))" + "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" + unfolding AQueue_def by simp_all + +text {* + \noindent For completeness, we provide a substitute for the + @{text case} combinator on queues: +*} + +definition %quote + aqueue_case_def: "aqueue_case = queue_case" + +lemma %quote aqueue_case [code, code inline]: + "queue_case = aqueue_case" + unfolding aqueue_case_def .. + +lemma %quote case_AQueue [code]: + "aqueue_case f (AQueue xs ys) = f (ys @ rev xs)" + unfolding aqueue_case_def AQueue_def by simp + +text {* + \noindent The resulting code looks as expected: +*} + +text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} + +text {* + \noindent From this example, it can be glimpsed that using own + constructor sets is a little delicate since it changes the set of + valid patterns for values of that type. Without going into much + detail, here some practical hints: + + \begin{itemize} + + \item When changing the constructor set for datatypes, take care + to provide an alternative for the @{text case} combinator + (e.g.~by replacing it using the preprocessor). + + \item Values in the target language need not to be normalised -- + different values in the target language may represent the same + value in the logic. + + \item Usually, a good methodology to deal with the subtleties of + pattern matching is to see the type as an abstract type: provide + a set of operations which operate on the concrete representation + of the type, and derive further operations by combinations of + these primitive ones, without relying on a particular + representation. + + \end{itemize} +*} + + +subsection {* Equality and wellsortedness *} + +text {* + Surely you have already noticed how equality is treated + by the code generator: +*} + +primrec %quote collect_duplicates :: "'a list \ 'a list \ 'a list \ 'a list" where + "collect_duplicates xs ys [] = xs" + | "collect_duplicates xs ys (z#zs) = (if z \ set xs + then if z \ set ys + then collect_duplicates xs ys zs + else collect_duplicates xs (z#ys) zs + else collect_duplicates (z#xs) (z#ys) zs)" + +text {* + \noindent The membership test during preprocessing is rewritten, + resulting in @{const List.member}, which itself + performs an explicit equality check. +*} + +text %quote {*@{code_stmts collect_duplicates (SML)}*} + +text {* + \noindent Obviously, polymorphic equality is implemented the Haskell + way using a type class. How is this achieved? HOL introduces + an explicit class @{class eq} with a corresponding operation + @{const eq_class.eq} such that @{thm eq [no_vars]}. + The preprocessing framework does the rest by propagating the + @{class eq} constraints through all dependent code equations. + For datatypes, instances of @{class eq} are implicitly derived + when possible. For other types, you may instantiate @{text eq} + manually like any other type class. + + Though this @{text eq} class is designed to get rarely in + the way, a subtlety + enters the stage when definitions of overloaded constants + are dependent on operational equality. For example, let + us define a lexicographic ordering on tuples + (also see theory @{theory Product_ord}): +*} + +instantiation %quote "*" :: (order, order) order +begin + +definition %quote [code del]: + "x \ y \ fst x < fst y \ fst x = fst y \ snd x \ snd y" + +definition %quote [code del]: + "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" + +instance %quote proof +qed (auto simp: less_eq_prod_def less_prod_def intro: order_less_trans) + +end %quote + +lemma %quote order_prod [code]: + "(x1 \ 'a\order, y1 \ 'b\order) < (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 < y2" + "(x1 \ 'a\order, y1 \ 'b\order) \ (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 \ y2" + by (simp_all add: less_prod_def less_eq_prod_def) + +text {* + \noindent Then code generation will fail. Why? The definition + of @{term "op \"} depends on equality on both arguments, + which are polymorphic and impose an additional @{class eq} + class constraint, which the preprocessor does not propagate + (for technical reasons). + + The solution is to add @{class eq} explicitly to the first sort arguments in the + code theorems: +*} + +lemma %quote order_prod_code [code]: + "(x1 \ 'a\{order, eq}, y1 \ 'b\order) < (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 < y2" + "(x1 \ 'a\{order, eq}, y1 \ 'b\order) \ (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 \ y2" + by (simp_all add: less_prod_def less_eq_prod_def) + +text {* + \noindent Then code generation succeeds: +*} + +text %quote {*@{code_stmts "op \ \ _ \ _ \ _ \ _ \ bool" (SML)}*} + +text {* + In some cases, the automatically derived code equations + for equality on a particular type may not be appropriate. + As example, watch the following datatype representing + monomorphic parametric types (where type constructors + are referred to by natural numbers): +*} + +datatype %quote monotype = Mono nat "monotype list" +(*<*) +lemma monotype_eq: + "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ + eq_class.eq tyco1 tyco2 \ eq_class.eq typargs1 typargs2" by (simp add: eq) +(*>*) + +text {* + \noindent Then code generation for SML would fail with a message + that the generated code contains illegal mutual dependencies: + the theorem @{thm monotype_eq [no_vars]} already requires the + instance @{text "monotype \ eq"}, which itself requires + @{thm monotype_eq [no_vars]}; Haskell has no problem with mutually + recursive @{text instance} and @{text function} definitions, + but the SML serialiser does not support this. + + In such cases, you have to provide your own equality equations + involving auxiliary constants. In our case, + @{const [show_types] list_all2} can do the job: +*} + +lemma %quote monotype_eq_list_all2 [code]: + "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ + eq_class.eq tyco1 tyco2 \ list_all2 eq_class.eq typargs1 typargs2" + by (simp add: eq list_all2_eq [symmetric]) + +text {* + \noindent does not depend on instance @{text "monotype \ eq"}: +*} + +text %quote {*@{code_stmts "eq_class.eq :: monotype \ monotype \ bool" (SML)}*} + + +subsection {* Explicit partiality *} + +text {* + Partiality usually enters the game by partial patterns, as + in the following example, again for amortised queues: +*} + +definition %quote strict_dequeue :: "'a queue \ 'a \ 'a queue" where + "strict_dequeue q = (case dequeue q + of (Some x, q') \ (x, q'))" + +lemma %quote strict_dequeue_AQueue [code]: + "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)" + "strict_dequeue (AQueue xs []) = + (case rev xs of y # ys \ (y, AQueue [] ys))" + by (simp_all add: strict_dequeue_def dequeue_AQueue split: list.splits) + +text {* + \noindent In the corresponding code, there is no equation + for the pattern @{term "AQueue [] []"}: +*} + +text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*} + +text {* + \noindent In some cases it is desirable to have this + pseudo-\qt{partiality} more explicitly, e.g.~as follows: +*} + +axiomatization %quote empty_queue :: 'a + +definition %quote strict_dequeue' :: "'a queue \ 'a \ 'a queue" where + "strict_dequeue' q = (case dequeue q of (Some x, q') \ (x, q') | _ \ empty_queue)" + +lemma %quote strict_dequeue'_AQueue [code]: + "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue + else strict_dequeue' (AQueue [] (rev xs)))" + "strict_dequeue' (AQueue xs (y # ys)) = + (y, AQueue xs ys)" + by (simp_all add: strict_dequeue'_def dequeue_AQueue split: list.splits) + +text {* + Observe that on the right hand side of the definition of @{const + "strict_dequeue'"} the constant @{const empty_queue} occurs + which is unspecified. + + Normally, if constants without any code equations occur in a + program, the code generator complains (since in most cases this is + not what the user expects). But such constants can also be thought + of as function definitions with no equations which always fail, + since there is never a successful pattern match on the left hand + side. In order to categorise a constant into that category + explicitly, use @{command "code_abort"}: +*} + +code_abort %quote empty_queue + +text {* + \noindent Then the code generator will just insert an error or + exception at the appropriate position: +*} + +text %quote {*@{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}*} + +text {* + \noindent This feature however is rarely needed in practice. + Note also that the @{text HOL} default setup already declares + @{const undefined} as @{command "code_abort"}, which is most + likely to be used in such situations. +*} + +end + \ No newline at end of file diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/ROOT.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,11 @@ + +(* $Id$ *) + +no_document use_thy "Setup"; +no_document use_thys ["Efficient_Nat"]; + +use_thy "Introduction"; +use_thy "Program"; +use_thy "Adaption"; +use_thy "Further"; +use_thy "ML"; diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/Setup.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Setup.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,12 @@ +theory Setup +imports Complex_Main +uses "../../../antiquote_setup.ML" "../../../more_antiquote.ML" +begin + +ML {* no_document use_thys + ["Efficient_Nat", "Code_Char_chr", "Product_ord", "~~/src/HOL/Imperative_HOL/Imperative_HOL", + "~~/src/HOL/Decision_Procs/Ferrack"] *} + +ML_val {* Code_Target.code_width := 74 *} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/document/Adaption.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Adaption.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,679 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Adaption}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Adaption\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isadeliminvisible +\isanewline +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{setup}\isamarkupfalse% +\ {\isacharverbatimopen}\ Code{\isacharunderscore}Target{\isachardot}extend{\isacharunderscore}target\ {\isacharparenleft}{\isachardoublequote}{\isasymSML}{\isachardoublequote}{\isacharcomma}\ {\isacharparenleft}{\isachardoublequote}SML{\isachardoublequote}{\isacharcomma}\ K\ I{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isamarkupsection{Adaption to target languages \label{sec:adaption}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Adapting code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The aspects of code generation introduced so far have two aspects + in common: + + \begin{itemize} + \item They act uniformly, without reference to a specific + target language. + \item They are \emph{safe} in the sense that as long as you trust + the code generator meta theory and implementation, you cannot + produce programs that yield results which are not derivable + in the logic. + \end{itemize} + + \noindent In this section we will introduce means to \emph{adapt} the serialiser + to a specific target language, i.e.~to print program fragments + in a way which accommodates \qt{already existing} ingredients of + a target language environment, for three reasons: + + \begin{itemize} + \item improving readability and aesthetics of generated code + \item gaining efficiency + \item interface with language parts which have no direct counterpart + in \isa{HOL} (say, imperative data structures) + \end{itemize} + + \noindent Generally, you should avoid using those features yourself + \emph{at any cost}: + + \begin{itemize} + \item The safe configuration methods act uniformly on every target language, + whereas for adaption you have to treat each target language separate. + \item Application is extremely tedious since there is no abstraction + which would allow for a static check, making it easy to produce garbage. + \item More or less subtle errors can be introduced unconsciously. + \end{itemize} + + \noindent However, even if you ought refrain from setting up adaption + yourself, already the \isa{HOL} comes with some reasonable default + adaptions (say, using target language list syntax). There also some + common adaption cases which you can setup by importing particular + library theories. In order to understand these, we provide some clues here; + these however are not supposed to replace a careful study of the sources.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{The adaption principle% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The following figure illustrates what \qt{adaption} is conceptually + supposed to be: + + \begin{figure}[here] + \begin{tikzpicture}[scale = 0.5] + \tikzstyle water=[color = blue, thick] + \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] + \tikzstyle process=[color = green, semithick, ->] + \tikzstyle adaption=[color = red, semithick, ->] + \tikzstyle target=[color = black] + \foreach \x in {0, ..., 24} + \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin + + (0.25, -0.25) cos + (0.25, 0.25); + \draw[style=ice] (1, 0) -- + (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; + \draw[style=ice] (9, 0) -- + (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; + \draw[style=ice] (15, -6) -- + (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; + \draw[style=process] + (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); + \draw[style=process] + (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); + \node (adaption) at (11, -2) [style=adaption] {adaption}; + \node at (19, 3) [rotate=90] {generated}; + \node at (19.5, -5) {language}; + \node at (19.5, -3) {library}; + \node (includes) at (19.5, -1) {includes}; + \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 + \draw[style=process] + (includes) -- (serialisation); + \draw[style=process] + (reserved) -- (serialisation); + \draw[style=adaption] + (adaption) -- (serialisation); + \draw[style=adaption] + (adaption) -- (includes); + \draw[style=adaption] + (adaption) -- (reserved); + \end{tikzpicture} + \caption{The adaption principle} + \label{fig:adaption} + \end{figure} + + \noindent In the tame view, code generation acts as broker between + \isa{logic}, \isa{intermediate\ language} and + \isa{target\ language} by means of \isa{translation} and + \isa{serialisation}; for the latter, the serialiser has to observe + the structure of the \isa{language} itself plus some \isa{reserved} + keywords which have to be avoided for generated code. + However, if you consider \isa{adaption} mechanisms, the code generated + by the serializer is just the tip of the iceberg: + + \begin{itemize} + \item \isa{serialisation} can be \emph{parametrised} such that + logical entities are mapped to target-specific ones + (e.g. target-specific list syntax, + see also \secref{sec:adaption_mechanisms}) + \item Such parametrisations can involve references to a + target-specific standard \isa{library} (e.g. using + the \isa{Haskell} \verb|Maybe| type instead + of the \isa{HOL} \isa{option} type); + if such are used, the corresponding identifiers + (in our example, \verb|Maybe|, \verb|Nothing| + and \verb|Just|) also have to be considered \isa{reserved}. + \item Even more, the user can enrich the library of the + target-language by providing code snippets + (\qt{\isa{includes}}) which are prepended to + any generated code (see \secref{sec:include}); this typically + also involves further \isa{reserved} identifiers. + \end{itemize} + + \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms + have to act consistently; it is at the discretion of the user + to take care for this.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Common adaption patterns% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \hyperlink{theory.Main}{\mbox{\isa{Main}}} theory already provides a code + generator setup + which should be suitable for most applications. Common extensions + and modifications are available by certain theories of the \isa{HOL} + library; beside being useful in applications, they may serve + as a tutorial for customising the code generator setup (see below + \secref{sec:adaption_mechanisms}). + + \begin{description} + + \item[\hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}}] represents \isa{HOL} integers by big + integer literals in target languages. + \item[\hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}] represents \isa{HOL} characters by + character literals in target languages. + \item[\hyperlink{theory.Code-Char-chr}{\mbox{\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}}}] like \isa{Code{\isacharunderscore}Char}, + but also offers treatment of character codes; includes + \hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}. + \item[\hyperlink{theory.Efficient-Nat}{\mbox{\isa{Efficient{\isacharunderscore}Nat}}}] \label{eff_nat} implements natural numbers by integers, + which in general will result in higher efficiency; pattern + matching with \isa{{\isadigit{0}}} / \isa{Suc} + is eliminated; includes \hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}} + and \hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}. + \item[\hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}] provides an additional datatype + \isa{index} which is mapped to target-language built-in integers. + Useful for code setups which involve e.g. indexing of + target-language arrays. + \item[\hyperlink{theory.Code-Message}{\mbox{\isa{Code{\isacharunderscore}Message}}}] provides an additional datatype + \isa{message{\isacharunderscore}string} which is isomorphic to strings; + \isa{message{\isacharunderscore}string}s are mapped to target-language strings. + Useful for code setups which involve e.g. printing (error) messages. + + \end{description} + + \begin{warn} + When importing any of these theories, they should form the last + items in an import list. Since these theories adapt the + code generator setup in a non-conservative fashion, + strange effects may occur otherwise. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Parametrising serialisation \label{sec:adaption_mechanisms}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Consider the following function and its corresponding + SML code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype boola = True | False;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun anda x True = x\\ +\hspace*{0pt} ~| anda x False = False\\ +\hspace*{0pt} ~| anda True x = x\\ +\hspace*{0pt} ~| anda False x = False;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = False\\ +\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Though this is correct code, it is a little bit unsatisfactory: + boolean values and operators are materialised as distinguished + entities with have nothing to do with the SML-built-in notion + of \qt{bool}. This results in less readable code; + additionally, eager evaluation may cause programs to + loop or break which would perfectly terminate when + the existing SML \verb|bool| would be used. To map + the HOL \isa{bool} on SML \verb|bool|, we may use + \qn{custom serialisations}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bool\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent The \hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} command takes a type constructor + as arguments together with a list of custom serialisations. + Each custom serialisation starts with a target language + identifier followed by an expression, which during + code serialisation is inserted whenever the type constructor + would occur. For constants, \hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} implements + the corresponding mechanism. Each ``\verb|_|'' in + a serialisation expression is treated as a placeholder + for the type constructor's (the constant's) arguments.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ +\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun in{\char95}interval (k,~l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This still is not perfect: the parentheses + around the \qt{andalso} expression are superfluous. + Though the serialiser + by no means attempts to imitate the rich Isabelle syntax + framework, it provides some common idioms, notably + associative infixes with precedences which may be used here:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ +\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun in{\char95}interval (k,~l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The attentive reader may ask how we assert that no generated + code will accidentally overwrite. For this reason the serialiser has + an internal table of identifiers which have to be avoided to be used + for new declarations. Initially, this table typically contains the + keywords of the target language. It can be extended manually, thus avoiding + accidental overwrites, using the \hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymSML}{\isachardoublequoteclose}\ bool\ true\ false\ andalso% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Next, we try to map HOL pairs to SML pairs, using the + infix ``\verb|*|'' type constructor and parentheses:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ {\isacharasterisk}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ Pair\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent The initial bang ``\verb|!|'' tells the serialiser + never to put + parentheses around the whole expression (they are already present), + while the parentheses around argument place holders + tell not to put parentheses around the arguments. + The slash ``\verb|/|'' (followed by arbitrary white space) + inserts a space which may be used as a break if necessary + during pretty printing. + + These examples give a glimpse what mechanisms + custom serialisations provide; however their usage + requires careful thinking in order not to introduce + inconsistencies -- or, in other words: + custom serialisations are completely axiomatic. + + A further noteworthy details is that any special + character in a custom serialisation may be quoted + using ``\verb|'|''; thus, in + ``\verb|fn '_ => _|'' the first + ``\verb|_|'' is a proper underscore while the + second ``\verb|_|'' is a placeholder.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{\isa{Haskell} serialisation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +For convenience, the default + \isa{HOL} setup for \isa{Haskell} maps the \isa{eq} class to + its counterpart in \isa{Haskell}, giving custom serialisations + for the class \isa{eq} (by command \hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}) and its operation + \isa{eq{\isacharunderscore}class{\isachardot}eq}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}class}\isamarkupfalse% +\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent A problem now occurs whenever a type which + is an instance of \isa{eq} in \isa{HOL} is mapped + on a \isa{Haskell}-built-in type which is also an instance + of \isa{Haskell} \isa{Eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{typedecl}\isamarkupfalse% +\ bar\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{by}\isamarkupfalse% +\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +\isanewline +% +\isadelimquotett +\isanewline +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bar\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent The code generator would produce + an additional instance, which of course is rejected by the \isa{Haskell} + compiler. + To suppress this additional instance, use + \isa{code{\isacharunderscore}instance}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\isamarkupsubsection{Enhancing the target language context \label{sec:include}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In rare cases it is necessary to \emph{enrich} the context of a + target language; this is accomplished using the \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} + command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}include}\isamarkupfalse% +\ Haskell\ {\isachardoublequoteopen}Errno{\isachardoublequoteclose}\isanewline +{\isacharverbatimopen}errno\ i\ {\isacharequal}\ error\ {\isacharparenleft}{\isachardoublequote}Error\ number{\isacharcolon}\ {\isachardoublequote}\ {\isacharplus}{\isacharplus}\ show\ i{\isacharparenright}{\isacharverbatimclose}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% +\ Haskell\ Errno% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent Such named \isa{include}s are then prepended to every generated code. + Inspect such code in order to find out how \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} behaves + with respect to a particular target language.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/document/Codegen.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Codegen.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,1690 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Codegen}% +% +\isadelimtheory +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isamarkupchapter{Code generation from Isabelle theories% +} +\isamarkuptrue% +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\isamarkupsubsection{Motivation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Executing formal specifications as programs is a well-established + topic in the theorem proving community. With increasing + application of theorem proving systems in the area of + software development and verification, its relevance manifests + for running test cases and rapid prototyping. In logical + calculi like constructive type theory, + a notion of executability is implicit due to the nature + of the calculus. In contrast, specifications in Isabelle + can be highly non-executable. In order to bridge + the gap between logic and executable specifications, + an explicit non-trivial transformation has to be applied: + code generation. + + This tutorial introduces a generic code generator for the + Isabelle system \cite{isa-tutorial}. + Generic in the sense that the + \qn{target language} for which code shall ultimately be + generated is not fixed but may be an arbitrary state-of-the-art + functional programming language (currently, the implementation + supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell + \cite{haskell-revised-report}). + We aim to provide a + versatile environment + suitable for software development and verification, + structuring the process + of code generation into a small set of orthogonal principles + while achieving a big coverage of application areas + with maximum flexibility. + + Conceptually the code generator framework is part + of Isabelle's \isa{Pure} meta logic; the object logic + \isa{HOL} which is an extension of \isa{Pure} + already comes with a reasonable framework setup and thus provides + a good working horse for raising code-generation-driven + applications. So, we assume some familiarity and experience + with the ingredients of the \isa{HOL} \emph{Main} theory + (see also \cite{isa-tutorial}).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Overview% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The code generator aims to be usable with no further ado + in most cases while allowing for detailed customization. + This manifests in the structure of this tutorial: + we start with a generic example \secref{sec:example} + and introduce code generation concepts \secref{sec:concept}. + Section + \secref{sec:basics} explains how to use the framework naively, + presuming a reasonable default setup. Then, section + \secref{sec:advanced} deals with advanced topics, + introducing further aspects of the code generator framework + in a motivation-driven manner. Last, section \secref{sec:ml} + introduces the framework's internal programming interfaces. + + \begin{warn} + Ultimately, the code generator which this tutorial deals with + is supposed to replace the already established code generator + by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. + So, for the moment, there are two distinct code generators + in Isabelle. + Also note that while the framework itself is + object-logic independent, only \isa{HOL} provides a reasonable + framework setup. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{An example: a simple theory of search trees \label{sec:example}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +When writing executable specifications using \isa{HOL}, + it is convenient to use + three existing packages: the datatype package for defining + datatypes, the function package for (recursive) functions, + and the class package for overloaded definitions. + + We develope a small theory of search trees; trees are represented + as a datatype with key type \isa{{\isacharprime}a} and value type \isa{{\isacharprime}b}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isacharequal}\ Leaf\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder{\isachardoublequoteclose}\ {\isacharprime}b\isanewline +\ \ {\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Note that we have constrained the type of keys + to the class of total orders, \isa{linorder}. + + We define \isa{find} and \isa{update} functions:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isasymColon}linorder{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}find\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isacharequal}\ key\ then\ Some\ val\ else\ None{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}find\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isasymle}\ key\ then\ find\ t{\isadigit{1}}\ it\ else\ find\ t{\isadigit{2}}\ it{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{fun}\isamarkupfalse% +\isanewline +\ \ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline +\ \ \ \ if\ it\ {\isacharequal}\ key\ then\ Leaf\ key\ entry\isanewline +\ \ \ \ \ \ else\ if\ it\ {\isasymle}\ key\isanewline +\ \ \ \ \ \ then\ Branch\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\isanewline +\ \ \ \ \ \ else\ Branch\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\isanewline +\ \ \ {\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline +\ \ \ \ if\ it\ {\isasymle}\ key\isanewline +\ \ \ \ \ \ then\ {\isacharparenleft}Branch\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{1}}{\isacharparenright}\ key\ t{\isadigit{2}}{\isacharparenright}\isanewline +\ \ \ \ \ \ else\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{2}}{\isacharparenright}{\isacharparenright}\isanewline +\ \ \ {\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent For testing purpose, we define a small example + using natural numbers \isa{nat} (which are a \isa{linorder}) + as keys and list of nats as values:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ example\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat{\isacharcomma}\ nat\ list{\isacharparenright}\ searchtree{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}Leaf\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Then we generate code% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ example\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}tree{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent which looks like: + \lstsml{Thy/examples/tree.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Code generation concepts and process \label{sec:concept}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{figure}[h] + \centering + \includegraphics[width=0.7\textwidth]{codegen_process} + \caption{code generator -- processing overview} + \label{fig:process} + \end{figure} + + The code generator employs a notion of executability + for three foundational executable ingredients known + from functional programming: + \emph{defining equations}, \emph{datatypes}, and + \emph{type classes}. A defining equation as a first approximation + is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} + (an equation headed by a constant \isa{f} with arguments + \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). + Code generation aims to turn defining equations + into a functional program by running through + a process (see figure \ref{fig:process}): + + \begin{itemize} + + \item Out of the vast collection of theorems proven in a + \qn{theory}, a reasonable subset modeling + defining equations is \qn{selected}. + + \item On those selected theorems, certain + transformations are carried out + (\qn{preprocessing}). Their purpose is to turn theorems + representing non- or badly executable + specifications into equivalent but executable counterparts. + The result is a structured collection of \qn{code theorems}. + + \item These \qn{code theorems} then are \qn{translated} + into an Haskell-like intermediate + language. + + \item Finally, out of the intermediate language the final + code in the desired \qn{target language} is \qn{serialized}. + + \end{itemize} + + From these steps, only the two last are carried out + outside the logic; by keeping this layer as + thin as possible, the amount of code to trust is + kept to a minimum.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Basics \label{sec:basics}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Invoking the code generator% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Thanks to a reasonable setup of the \isa{HOL} theories, in + most cases code generation proceeds without further ado:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ fac\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}fac\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}fac\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharasterisk}\ fac\ n{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent This executable specification is now turned to SML code:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ fac\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fac{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent The \isa{{\isasymEXPORTCODE}} command takes a space-separated list of + constants together with \qn{serialization directives} + These start with a \qn{target language} + identifier, followed by a file specification + where to write the generated code to. + + Internally, the defining equations for all selected + constants are taken, including any transitively required + constants, datatypes and classes, resulting in the following + code: + + \lstsml{Thy/examples/fac.ML} + + The code generator will complain when a required + ingredient does not provide a executable counterpart, + e.g.~generating code + for constants not yielding + a defining equation (e.g.~the Hilbert choice + operation \isa{SOME}):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ pick{\isacharunderscore}some\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pick{\isacharunderscore}some\ xs\ {\isacharequal}\ {\isacharparenleft}SOME\ x{\isachardot}\ x\ {\isasymin}\ set\ xs{\isacharparenright}{\isachardoublequoteclose}% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isanewline +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ pick{\isacharunderscore}some\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fail{\isacharunderscore}const{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent will fail.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Theorem selection% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The list of all defining equations in a theory may be inspected + using the \isa{{\isasymPRINTCODESETUP}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{print{\isacharunderscore}codesetup}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent which displays a table of constant with corresponding + defining equations (the additional stuff displayed + shall not bother us for the moment). + + The typical \isa{HOL} tools are already set up in a way that + function definitions introduced by \isa{{\isasymDEFINITION}}, + \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, + \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, + \isa{{\isasymRECDEF}} are implicitly propagated + to this defining equation table. Specific theorems may be + selected using an attribute: \emph{code func}. As example, + a weight selector function:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}let\ {\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}\ {\isacharequal}\ x\ in\isanewline +\ \ \ \ if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We want to eliminate the explicit destruction + of \isa{x} to \isa{{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ pick\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}pick{\isadigit{1}}{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent This theorem now is used for generating code: + + \lstsml{Thy/examples/pick1.ML} + + \noindent The policy is that \emph{default equations} stemming from + \isa{{\isasymDEFINITION}}, + \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, + \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, + \isa{{\isasymRECDEF}} statements are discarded as soon as an + equation is explicitly selected by means of \emph{code func}. + Further applications of \emph{code func} add theorems incrementally, + but syntactic redundancies are implicitly dropped. For example, + using a modified version of the \isa{fac} function + as defining equation, the then redundant (since + syntactically subsumed) original defining equations + are dropped. + + \begin{warn} + The attributes \emph{code} and \emph{code del} + associated with the existing code generator also apply to + the new one: \emph{code} implies \emph{code func}, + and \emph{code del} implies \emph{code func del}. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Type classes% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Type classes enter the game via the Isar class package. + For a short introduction how to use it, see \cite{isabelle-classes}; + here we just illustrate its impact on code generation. + + In a target language, type classes may be represented + natively (as in the case of Haskell). For languages + like SML, they are implemented using \emph{dictionaries}. + Our following example specifies a class \qt{null}, + assigning to each of its inhabitants a \qt{null} value:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{class}\isamarkupfalse% +\ null\ {\isacharequal}\ type\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ null\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}head\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ null{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}head\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We provide some instances for our \isa{null}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{instantiation}\isamarkupfalse% +\ option\ \isakeyword{and}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ null\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ None{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Constructing a dummy example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isachardoublequoteopen}dummy\ {\isacharequal}\ head\ {\isacharbrackleft}Some\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ None{\isacharbrackright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Type classes offer a suitable occasion to introduce + the Haskell serializer. Its usage is almost the same + as SML, but, in accordance with conventions + some Haskell systems enforce, each module ends + up in a single file. The module hierarchy is reflected in + the file system, with root directory given as file specification.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ dummy\ \isakeyword{in}\ Haskell\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lsthaskell{Thy/examples/Codegen.hs} + \noindent (we have left out all other modules). + + \medskip + + The whole code in SML with explicit dictionary passing:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ dummy\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/class.ML} + + \medskip + + \noindent or in OCaml:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ dummy\ \isakeyword{in}\ OCaml\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ocaml{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/class.ocaml} + + \medskip The explicit association of constants + to classes can be inspected using the \isa{{\isasymPRINTCLASSES}} + command.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Recipes and advanced topics \label{sec:advanced}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In this tutorial, we do not attempt to give an exhaustive + description of the code generator framework; instead, + we cast a light on advanced topics by introducing + them together with practically motivated examples. Concerning + further reading, see + + \begin{itemize} + + \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} + for exhaustive syntax diagrams. + \item or \cite{Haftmann-Nipkow:2007:codegen} which deals with foundational issues + of the code generator framework. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Library theories \label{sec:library}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \isa{HOL} \isa{Main} theory already provides a code + generator setup + which should be suitable for most applications. Common extensions + and modifications are available by certain theories of the \isa{HOL} + library; beside being useful in applications, they may serve + as a tutorial for customizing the code generator setup. + + \begin{description} + + \item[\isa{Code{\isacharunderscore}Integer}] represents \isa{HOL} integers by big + integer literals in target languages. + \item[\isa{Code{\isacharunderscore}Char}] represents \isa{HOL} characters by + character literals in target languages. + \item[\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}] like \isa{Code{\isacharunderscore}Char}, + but also offers treatment of character codes; includes + \isa{Code{\isacharunderscore}Integer}. + \item[\isa{Efficient{\isacharunderscore}Nat}] \label{eff_nat} implements natural numbers by integers, + which in general will result in higher efficency; pattern + matching with \isa{{\isadigit{0}}} / \isa{Suc} + is eliminated; includes \isa{Code{\isacharunderscore}Integer}. + \item[\isa{Code{\isacharunderscore}Index}] provides an additional datatype + \isa{index} which is mapped to target-language built-in integers. + Useful for code setups which involve e.g. indexing of + target-language arrays. + \item[\isa{Code{\isacharunderscore}Message}] provides an additional datatype + \isa{message{\isacharunderscore}string} which is isomorphic to strings; + \isa{message{\isacharunderscore}string}s are mapped to target-language strings. + Useful for code setups which involve e.g. printing (error) messages. + + \end{description} + + \begin{warn} + When importing any of these theories, they should form the last + items in an import list. Since these theories adapt the + code generator setup in a non-conservative fashion, + strange effects may occur otherwise. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Preprocessing% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Before selected function theorems are turned into abstract + code, a chain of definitional transformation steps is carried + out: \emph{preprocessing}. In essence, the preprocessor + consists of two components: a \emph{simpset} and \emph{function transformers}. + + The \emph{simpset} allows to employ the full generality of the Isabelle + simplifier. Due to the interpretation of theorems + as defining equations, rewrites are applied to the right + hand side and the arguments of the left hand side of an + equation, but never to the constant heading the left hand side. + An important special case are \emph{inline theorems} which may be + declared an undeclared using the + \emph{code inline} or \emph{code inline del} attribute respectively. + Some common applications:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{itemize} +% +\begin{isamarkuptext}% +\item replacing non-executable constructs by executable ones:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\item eliminating superfluous constants:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\item replacing executable but inconvenient constructs:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{itemize} +% +\begin{isamarkuptext}% +\emph{Function transformers} provide a very general interface, + transforming a list of function theorems to another + list of function theorems, provided that neither the heading + constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} + pattern elimination implemented in + theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this + interface. + + \noindent The current setup of the preprocessor may be inspected using + the \isa{{\isasymPRINTCODESETUP}} command. + + \begin{warn} + The attribute \emph{code unfold} + associated with the existing code generator also applies to + the new one: \emph{code unfold} implies \emph{code inline}. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Concerning operational equality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Surely you have already noticed how equality is treated + by the code generator:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline +\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline +\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline +\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline +\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +The membership test during preprocessing is rewritten, + resulting in \isa{op\ mem}, which itself + performs an explicit equality check.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ collect{\isacharunderscore}duplicates\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}collect{\isacharunderscore}duplicates{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/collect_duplicates.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Obviously, polymorphic equality is implemented the Haskell + way using a type class. How is this achieved? HOL introduces + an explicit class \isa{eq} with a corresponding operation + \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}}. + The preprocessing framework does the rest. + For datatypes, instances of \isa{eq} are implicitly derived + when possible. For other types, you may instantiate \isa{eq} + manually like any other type class. + + Though this \isa{eq} class is designed to get rarely in + the way, a subtlety + enters the stage when definitions of overloaded constants + are dependent on operational equality. For example, let + us define a lexicographic ordering on tuples:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ord{\isacharcomma}\ ord{\isacharparenright}\ ord\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ ord{\isacharunderscore}prod\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Then code generation will fail. Why? The definition + of \isa{op\ {\isasymle}} depends on equality on both arguments, + which are polymorphic and impose an additional \isa{eq} + class constraint, thus violating the type discipline + for class operations. + + The solution is to add \isa{eq} explicitly to the first sort arguments in the + code theorems:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ ord{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ ord{\isacharunderscore}prod\ \isacommand{by}\isamarkupfalse% +\ rule{\isacharplus}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Then code generation succeeds:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}ord\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}lexicographic{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/lexicographic.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +In general, code theorems for overloaded constants may have more + restrictive sort constraints than the underlying instance relation + between class and type constructor as long as the whole system of + constraints is coregular; code theorems violating coregularity + are rejected immediately. Consequently, it might be necessary + to delete disturbing theorems in the code theorem table, + as we have done here with the original definitions \isa{less{\isacharunderscore}prod{\isacharunderscore}def} + and \isa{less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def}. + + In some cases, the automatically derived defining equations + for equality on a particular type may not be appropriate. + As example, watch the following datatype representing + monomorphic parametric types (where type constructors + are referred to by natural numbers):% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Then code generation for SML would fail with a message + that the generated code conains illegal mutual dependencies: + the theorem \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}} already requires the + instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires + \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually + recursive \isa{instance} and \isa{function} definitions, + but the SML serializer does not support this. + + In such cases, you have to provide you own equality equations + involving auxiliary constants. In our case, + \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ {\isacharparenleft}op\ {\isacharequal}{\isacharparenright}\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharequal}\ {\isacharcolon}{\isacharcolon}\ monotype\ {\isasymRightarrow}\ monotype\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}monotype{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/monotype.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Programs as sets of theorems% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As told in \secref{sec:concept}, code generation is based + on a structured collection of code theorems. + For explorative purpose, this collection + may be inspected using the \isa{{\isasymCODETHMS}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ mod\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent prints a table with \emph{all} defining equations + for \isa{op\ mod}, including + \emph{all} defining equations those equations depend + on recursivly. \isa{{\isasymCODETHMS}} provides a convenient + mechanism to inspect the impact of a preprocessor setup + on defining equations. + + Similarly, the \isa{{\isasymCODEDEPS}} command shows a graph + visualizing dependencies between defining equations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Constructor sets for datatypes% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Conceptually, any datatype is spanned by a set of + \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} + where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is excactly the set of \emph{all} + type variables in \isa{{\isasymtau}}. The HOL datatype package + by default registers any new datatype in the table + of datatypes, which may be inspected using + the \isa{{\isasymPRINTCODESETUP}} command. + + In some cases, it may be convenient to alter or + extend this table; as an example, we will develope an alternative + representation of natural numbers as binary digits, whose + size does increase logarithmically with its value, not linear + \footnote{Indeed, the \isa{Efficient{\isacharunderscore}Nat} theory (see \ref{eff_nat}) + does something similar}. First, the digit representation:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\ Dig{\isadigit{0}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ n{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ Dig{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We will use these two ">digits"< to represent natural numbers + in binary digits, e.g.:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isadigit{4}}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Of course we also have to provide proper code equations for + the operations, e.g. \isa{op\ {\isacharplus}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}m\ {\isacharplus}\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ n{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{1}}\ m{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent We then instruct the code generator to view \isa{{\isadigit{0}}}, + \isa{{\isadigit{1}}}, \isa{Dig{\isadigit{0}}} and \isa{Dig{\isadigit{1}}} as + datatype constructors:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{0}}{\isasymColon}nat{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isadigit{1}}{\isasymColon}nat{\isachardoublequoteclose}\ Dig{\isadigit{0}}\ Dig{\isadigit{1}}% +\begin{isamarkuptext}% +\noindent For the former constructor \isa{Suc}, we provide a code + equation and remove some parts of the default code generator setup + which are an obstacle here:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ Suc{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}Suc\ n\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{declare}\isamarkupfalse% +\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline\ del{\isacharbrackright}\isanewline +\isacommand{declare}\isamarkupfalse% +\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}% +\begin{isamarkuptext}% +\noindent This yields the following code:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharplus}\ {\isasymColon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}nat{\isacharunderscore}binary{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/nat_binary.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip + + From this example, it can be easily glimpsed that using own constructor sets + is a little delicate since it changes the set of valid patterns for values + of that type. Without going into much detail, here some practical hints: + + \begin{itemize} + \item When changing the constuctor set for datatypes, take care to + provide an alternative for the \isa{case} combinator (e.g. by replacing + it using the preprocessor). + \item Values in the target language need not to be normalized -- different + values in the target language may represent the same value in the + logic (e.g. \isa{Dig{\isadigit{1}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}}). + \item Usually, a good methodology to deal with the subleties of pattern + matching is to see the type as an abstract type: provide a set + of operations which operate on the concrete representation of the type, + and derive further operations by combinations of these primitive ones, + without relying on a particular representation. + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isachardoublequoteclose}\ Suc\isanewline +\isacommand{declare}\isamarkupfalse% +\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}\isanewline +\isacommand{declare}\isamarkupfalse% +\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline{\isacharbrackright}\isanewline +\isacommand{declare}\isamarkupfalse% +\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func{\isacharbrackright}\ \isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymColon}\ nat{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isamarkupsubsection{Customizing serialization% +} +\isamarkuptrue% +% +\isamarkupsubsubsection{Basics% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Consider the following function and its corresponding + SML code:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ in{\isacharunderscore}interval\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}literal{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/bool_literal.ML} + + \noindent Though this is correct code, it is a little bit unsatisfactory: + boolean values and operators are materialized as distinguished + entities with have nothing to do with the SML-builtin notion + of \qt{bool}. This results in less readable code; + additionally, eager evaluation may cause programs to + loop or break which would perfectly terminate when + the existing SML \qt{bool} would be used. To map + the HOL \qt{bool} on SML \qt{bool}, we may use + \qn{custom serializations}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bool\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +The \isa{{\isasymCODETYPE}} commad takes a type constructor + as arguments together with a list of custom serializations. + Each custom serialization starts with a target language + identifier followed by an expression, which during + code serialization is inserted whenever the type constructor + would occur. For constants, \isa{{\isasymCODECONST}} implements + the corresponding mechanism. Each ``\verb|_|'' in + a serialization expression is treated as a placeholder + for the type constructor's (the constant's) arguments.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}mlbool{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/bool_mlbool.ML} + + \noindent This still is not perfect: the parentheses + around the \qt{andalso} expression are superfluous. + Though the serializer + by no means attempts to imitate the rich Isabelle syntax + framework, it provides some common idioms, notably + associative infixes with precedences which may be used here:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +\isanewline +\isanewline +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}infix{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/bool_infix.ML} + + \medskip + + Next, we try to map HOL pairs to SML pairs, using the + infix ``\verb|*|'' type constructor and parentheses:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ {\isacharasterisk}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ Pair\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +The initial bang ``\verb|!|'' tells the serializer to never put + parentheses around the whole expression (they are already present), + while the parentheses around argument place holders + tell not to put parentheses around the arguments. + The slash ``\verb|/|'' (followed by arbitrary white space) + inserts a space which may be used as a break if necessary + during pretty printing. + + These examples give a glimpse what mechanisms + custom serializations provide; however their usage + requires careful thinking in order not to introduce + inconsistencies -- or, in other words: + custom serializations are completely axiomatic. + + A further noteworthy details is that any special + character in a custom serialization may be quoted + using ``\verb|'|''; thus, in + ``\verb|fn '_ => _|'' the first + ``\verb|_|'' is a proper underscore while the + second ``\verb|_|'' is a placeholder. + + The HOL theories provide further + examples for custom serializations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Haskell serialization% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +For convenience, the default + HOL setup for Haskell maps the \isa{eq} class to + its counterpart in Haskell, giving custom serializations + for the class (\isa{{\isasymCODECLASS}}) and its operation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}class}\isamarkupfalse% +\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}\ \isakeyword{where}\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharequal}{\isacharequal}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +A problem now occurs whenever a type which + is an instance of \isa{eq} in HOL is mapped + on a Haskell-builtin type which is also an instance + of Haskell \isa{Eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{typedecl}\isamarkupfalse% +\ bar\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +% +\isadelimtt +\isanewline +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bar\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +The code generator would produce + an additional instance, which of course is rejected. + To suppress this additional instance, use + \isa{{\isasymCODEINSTANCE}}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\isamarkupsubsubsection{Pretty printing% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The serializer provides ML interfaces to set up + pretty serializations for expressions like lists, numerals + and characters; these are + monolithic stubs and should only be used with the + theories introduced in \secref{sec:library}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Cyclic module dependencies% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Sometimes the awkward situation occurs that dependencies + between definitions introduce cyclic dependencies + between modules, which in the Haskell world leaves + you to the mercy of the Haskell implementation you are using, + while for SML code generation is not possible. + + A solution is to declare module names explicitly. + Let use assume the three cyclically dependent + modules are named \emph{A}, \emph{B} and \emph{C}. + Then, by stating% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% +\ SML\isanewline +\ \ A\ ABC\isanewline +\ \ B\ ABC\isanewline +\ \ C\ ABC% +\begin{isamarkuptext}% +we explicitly map all those modules on \emph{ABC}, + resulting in an ad-hoc merge of this three modules + at serialization time.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Incremental code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Code generation is \emph{incremental}: theorems + and abstract intermediate code are cached and extended on demand. + The cache may be partially or fully dropped if the underlying + executable content of the theory changes. + Implementation of caching is supposed to transparently + hid away the details from the user. Anyway, caching + reaches the surface by using a slightly more general form + of the \isa{{\isasymCODETHMS}}, \isa{{\isasymCODEDEPS}} + and \isa{{\isasymEXPORTCODE}} commands: the list of constants + may be omitted. Then, all constants with code theorems + in the current cache are referred to.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{ML interfaces \label{sec:ml}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Since the code generator framework not only aims to provide + a nice Isar interface but also to form a base for + code-generation-based applications, here a short + description of the most important ML interfaces.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Executable theory content: \isa{Code}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This Pure module implements the core notions of + executable content of a theory.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Managing executable content% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexml{Code.add\_func}\verb|Code.add_func: thm -> theory -> theory| \\ + \indexml{Code.del\_func}\verb|Code.del_func: thm -> theory -> theory| \\ + \indexml{Code.add\_funcl}\verb|Code.add_funcl: string * thm list Susp.T -> theory -> theory| \\ + \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ + \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ + \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> thm list -> thm list option)|\isasep\isanewline% +\verb| -> theory -> theory| \\ + \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ + \indexml{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ + \indexml{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% +\verb| -> (string * sort) list * (string * typ list) list| \\ + \indexml{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| + \end{mldecls} + + \begin{description} + + \item \verb|Code.add_func|~\isa{thm}~\isa{thy} adds function + theorem \isa{thm} to executable content. + + \item \verb|Code.del_func|~\isa{thm}~\isa{thy} removes function + theorem \isa{thm} from executable content, if present. + + \item \verb|Code.add_funcl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds + suspended defining equations \isa{lthms} for constant + \isa{const} to executable content. + + \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes + the preprocessor simpset. + + \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds + function transformer \isa{f} (named \isa{name}) to executable content; + \isa{f} is a transformer of the defining equations belonging + to a certain function definition, depending on the + current theory context. Returning \isa{NONE} indicates that no + transformation took place; otherwise, the whole process will be iterated + with the new defining equations. + + \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes + function transformer named \isa{name} from executable content. + + \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds + a datatype to executable content, with generation + set \isa{cs}. + + \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} + returns type constructor corresponding to + constructor \isa{const}; returns \isa{NONE} + if \isa{const} is no constructor. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Auxiliary% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ + \indexml{Code\_Unit.head\_func}\verb|Code_Unit.head_func: thm -> string * ((string * sort) list * typ)| \\ + \indexml{Code\_Unit.rewrite\_func}\verb|Code_Unit.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\ + \end{mldecls} + + \begin{description} + + \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} + reads a constant as a concrete term expression \isa{s}. + + \item \verb|Code_Unit.head_func|~\isa{thm} + extracts the constant and its type from a defining equation \isa{thm}. + + \item \verb|Code_Unit.rewrite_func|~\isa{ss}~\isa{thm} + rewrites a defining equation \isa{thm} with a simpset \isa{ss}; + only arguments and right hand side are rewritten, + not the head of the defining equation. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Implementing code generator applications% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Implementing code generator applications on top + of the framework set out so far usually not only + involves using those primitive interfaces + but also storing code-dependent data and various + other things. + + \begin{warn} + Some interfaces discussed here have not reached + a final state yet. + Changes likely to occur in future. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Data depending on the theory's executable content% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Due to incrementality of code generation, changes in the + theory's executable content have to be propagated in a + certain fashion. Additionally, such changes may occur + not only during theory extension but also during theory + merge, which is a little bit nasty from an implementation + point of view. The framework provides a solution + to this technical challenge by providing a functorial + data slot \verb|CodeDataFun|; on instantiation + of this functor, the following types and operations + are required: + + \medskip + \begin{tabular}{l} + \isa{type\ T} \\ + \isa{val\ empty{\isacharcolon}\ T} \\ + \isa{val\ merge{\isacharcolon}\ Pretty{\isachardot}pp\ {\isasymrightarrow}\ T\ {\isacharasterisk}\ T\ {\isasymrightarrow}\ T} \\ + \isa{val\ purge{\isacharcolon}\ theory\ option\ {\isasymrightarrow}\ CodeUnit{\isachardot}const\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} + \end{tabular} + + \begin{description} + + \item \isa{T} the type of data to store. + + \item \isa{empty} initial (empty) data. + + \item \isa{merge} merging two data slots. + + \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; + if possible, the current theory context is handed over + as argument \isa{thy} (if there is no current theory context (e.g.~during + theory merge, \verb|NONE|); \isa{consts} indicates the kind + of change: \verb|NONE| stands for a fundamental change + which invalidates any existing code, \isa{SOME\ consts} + hints that executable content for constants \isa{consts} + has changed. + + \end{description} + + An instance of \verb|CodeDataFun| provides the following + interface: + + \medskip + \begin{tabular}{l} + \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} + \end{tabular} + + \begin{description} + + \item \isa{get} retrieval of the current data. + + \item \isa{change} update of current data (cached!) + by giving a continuation. + + \item \isa{change{\isacharunderscore}yield} update with side result. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\emph{Happy proving, happy hacking!}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/document/Further.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Further.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,234 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Further}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Further\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Further issues \label{sec:further}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Further reading% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Do dive deeper into the issue of code generation, you should visit + the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which + contains exhaustive syntax diagrams.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Modules% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +When invoking the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command it is possible to leave + out the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isacharunderscore}name}}}} part; then code is distributed over + different modules, where the module name space roughly is induced + by the \isa{Isabelle} theory name space. + + Then sometimes the awkward situation occurs that dependencies between + definitions introduce cyclic dependencies between modules, which in the + \isa{Haskell} world leaves you to the mercy of the \isa{Haskell} implementation + you are using, while for \isa{SML}/\isa{OCaml} code generation is not possible. + + A solution is to declare module names explicitly. + Let use assume the three cyclically dependent + modules are named \emph{A}, \emph{B} and \emph{C}. + Then, by stating% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% +\ SML\isanewline +\ \ A\ ABC\isanewline +\ \ B\ ABC\isanewline +\ \ C\ ABC% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +we explicitly map all those modules on \emph{ABC}, + resulting in an ad-hoc merge of this three modules + at serialisation time.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Evaluation oracle% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Code generation may also be used to \emph{evaluate} expressions + (using \isa{SML} as target language of course). + For instance, the \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} allows to reduce an expression to a + normal form with respect to the underlying code equations:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{value}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent will display \isa{{\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}}. + + The \hyperlink{method.eval}{\mbox{\isa{eval}}} method tries to reduce a goal by code generation to \isa{True} + and solves it in that case, but fails otherwise:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}\ {\isacharequal}\ {\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ eval% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The soundness of the \hyperlink{method.eval}{\mbox{\isa{eval}}} method depends crucially + on the correctness of the code generator; this is one of the reasons + why you should not use adaption (see \secref{sec:adaption}) frivolously.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Code antiquotation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In scenarios involving techniques like reflection it is quite common + that code generated from a theory forms the basis for implementing + a proof procedure in \isa{SML}. To facilitate interfacing of generated code + with system code, the code generator provides a \isa{code} antiquotation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ form\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ And\ form\ form\ {\isacharbar}\ Or\ form\ form\isanewline +\isanewline +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ fun\ eval{\isacharunderscore}form\ % +\isaantiq +code\ T% +\endisaantiq +\ {\isacharequal}\ true\isanewline +\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ % +\isaantiq +code\ F% +\endisaantiq +\ {\isacharequal}\ false\isanewline +\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% +\isaantiq +code\ And% +\endisaantiq +\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ andalso\ eval{\isacharunderscore}form\ q\isanewline +\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% +\isaantiq +code\ Or% +\endisaantiq +\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ orelse\ eval{\isacharunderscore}form\ q{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent \isa{code} takes as argument the name of a constant; after the + whole \isa{SML} is read, the necessary code is generated transparently + and the corresponding constant names are inserted. This technique also + allows to use pattern matching on constructors stemming from compiled + \isa{datatypes}. + + For a less simplistic example, theory \hyperlink{theory.Ferrack}{\mbox{\isa{Ferrack}}} is + a good reference.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Imperative data structures% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +If you consider imperative data structures as inevitable for a specific + application, you should consider + \emph{Imperative Functional Programming with Isabelle/HOL} + (\cite{bulwahn-et-al:2008:imperative}); + the framework described there is available in theory \hyperlink{theory.Imperative-HOL}{\mbox{\isa{Imperative{\isacharunderscore}HOL}}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/document/Introduction.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Introduction.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,390 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Introduction}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Introduction\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Code generation from \isa{Isabelle{\isacharslash}HOL} theories% +} +\isamarkuptrue% +% +\isamarkupsection{Introduction and Overview% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This tutorial introduces a generic code generator for the + \isa{Isabelle} system. + Generic in the sense that the + \qn{target language} for which code shall ultimately be + generated is not fixed but may be an arbitrary state-of-the-art + functional programming language (currently, the implementation + supports \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and \isa{Haskell} + \cite{haskell-revised-report}). + + Conceptually the code generator framework is part + of Isabelle's \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} meta logic framework; the logic + \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} which is an extension of \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} + already comes with a reasonable framework setup and thus provides + a good working horse for raising code-generation-driven + applications. So, we assume some familiarity and experience + with the ingredients of the \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} distribution theories. + (see also \cite{isa-tutorial}). + + The code generator aims to be usable with no further ado + in most cases while allowing for detailed customisation. + This manifests in the structure of this tutorial: after a short + conceptual introduction with an example (\secref{sec:intro}), + we discuss the generic customisation facilities (\secref{sec:program}). + A further section (\secref{sec:adaption}) is dedicated to the matter of + \qn{adaption} to specific target language environments. After some + further issues (\secref{sec:further}) we conclude with an overview + of some ML programming interfaces (\secref{sec:ml}). + + \begin{warn} + Ultimately, the code generator which this tutorial deals with + is supposed to replace the existing code generator + by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. + So, for the moment, there are two distinct code generators + in Isabelle. In case of ambiguity, we will refer to the framework + described here as \isa{generic\ code\ generator}, to the + other as \isa{SML\ code\ generator}. + Also note that while the framework itself is + object-logic independent, only \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} provides a reasonable + framework setup. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Code generation via shallow embedding \label{sec:intro}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The key concept for understanding \isa{Isabelle}'s code generation is + \emph{shallow embedding}, i.e.~logical entities like constants, types and + classes are identified with corresponding concepts in the target language. + + Inside \hyperlink{theory.HOL}{\mbox{\isa{HOL}}}, the \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} and + \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} declarations form + the core of a functional programming language. The default code generator setup + allows to turn those into functional programs immediately. + This means that \qt{naive} code generation can proceed without further ado. + For example, here a simple \qt{implementation} of amortised queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ {\isacharprime}a\ queue\ {\isacharequal}\ AQueue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{fun}\isamarkupfalse% +\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then we can generate code e.g.~for \isa{SML} as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ empty\ dequeue\ enqueue\ \isakeyword{in}\ SML\isanewline +\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}example{\isachardot}ML{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent resulting in the following code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun foldl f a [] = a\\ +\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun list{\char95}case f1 f2 (a ::~lista) = f2 a lista\\ +\hspace*{0pt} ~| list{\char95}case f1 f2 [] = f1;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun dequeue (AQueue ([],~[])) = (NONE,~AQueue ([],~[]))\\ +\hspace*{0pt} ~| dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ +\hspace*{0pt} ~| dequeue (AQueue (v ::~va,~[])) =\\ +\hspace*{0pt} ~~~let\\ +\hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\ +\hspace*{0pt} ~~~in\\ +\hspace*{0pt} ~~~~~(SOME y,~AQueue ([],~ys))\\ +\hspace*{0pt} ~~~end;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command takes a space-separated list of + constants for which code shall be generated; anything else needed for those + is added implicitly. Then follows a target language identifier + (\isa{SML}, \isa{OCaml} or \isa{Haskell}) and a freely chosen module name. + A file name denotes the destination to store the generated code. Note that + the semantics of the destination depends on the target language: for + \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell} + it denotes a \emph{directory} where a file named as the module name + (with extension \isa{{\isachardot}hs}) is written:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ empty\ dequeue\ enqueue\ \isakeyword{in}\ Haskell\isanewline +\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This is how the corresponding code in \isa{Haskell} looks like:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}module Example where {\char123}\\ +\hspace*{0pt}\\ +\hspace*{0pt}\\ +\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\ +\hspace*{0pt}foldla f a [] = a;\\ +\hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\ +\hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}list{\char95}case ::~forall t a.~t -> (a -> [a] -> t) -> [a] -> t;\\ +\hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\ +\hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\ +\hspace*{0pt}\\ +\hspace*{0pt}data Queue a = AQueue [a] [a];\\ +\hspace*{0pt}\\ +\hspace*{0pt}empty ::~forall a.~Queue a;\\ +\hspace*{0pt}empty = AQueue [] [];\\ +\hspace*{0pt}\\ +\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ +\hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\ +\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ +\hspace*{0pt}dequeue (AQueue (v :~va) []) =\\ +\hspace*{0pt} ~let {\char123}\\ +\hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\ +\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\ +\hspace*{0pt}\\ +\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\ +\hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\ +\hspace*{0pt}\\ +\hspace*{0pt}{\char125}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This demonstrates the basic usage of the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command; + for more details see \secref{sec:further}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Code generator architecture \label{sec:concept}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +What you have seen so far should be already enough in a lot of cases. If you + are content with this, you can quit reading here. Anyway, in order to customise + and adapt the code generator, it is inevitable to gain some understanding + how it works. + + \begin{figure}[h] + \begin{tikzpicture}[x = 4.2cm, y = 1cm] + \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; + \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; + \tikzstyle process_arrow=[->, semithick, color = green]; + \node (HOL) at (0, 4) [style=entity] {\isa{Isabelle{\isacharslash}HOL} theory}; + \node (eqn) at (2, 2) [style=entity] {code equations}; + \node (iml) at (2, 0) [style=entity] {intermediate language}; + \node (seri) at (1, 0) [style=process] {serialisation}; + \node (SML) at (0, 3) [style=entity] {\isa{SML}}; + \node (OCaml) at (0, 2) [style=entity] {\isa{OCaml}}; + \node (further) at (0, 1) [style=entity] {\isa{{\isasymdots}}}; + \node (Haskell) at (0, 0) [style=entity] {\isa{Haskell}}; + \draw [style=process_arrow] (HOL) .. controls (2, 4) .. + node [style=process, near start] {selection} + node [style=process, near end] {preprocessing} + (eqn); + \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); + \draw [style=process_arrow] (iml) -- (seri); + \draw [style=process_arrow] (seri) -- (SML); + \draw [style=process_arrow] (seri) -- (OCaml); + \draw [style=process_arrow, dashed] (seri) -- (further); + \draw [style=process_arrow] (seri) -- (Haskell); + \end{tikzpicture} + \caption{Code generator architecture} + \label{fig:arch} + \end{figure} + + The code generator employs a notion of executability + for three foundational executable ingredients known + from functional programming: + \emph{code equations}, \emph{datatypes}, and + \emph{type classes}. A code equation as a first approximation + is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} + (an equation headed by a constant \isa{f} with arguments + \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). + Code generation aims to turn code equations + into a functional program. This is achieved by three major + components which operate sequentially, i.e. the result of one is + the input + of the next in the chain, see diagram \ref{fig:arch}: + + \begin{itemize} + + \item Out of the vast collection of theorems proven in a + \qn{theory}, a reasonable subset modelling + code equations is \qn{selected}. + + \item On those selected theorems, certain + transformations are carried out + (\qn{preprocessing}). Their purpose is to turn theorems + representing non- or badly executable + specifications into equivalent but executable counterparts. + The result is a structured collection of \qn{code theorems}. + + \item Before the selected code equations are continued with, + they can be \qn{preprocessed}, i.e. subjected to theorem + transformations. This \qn{preprocessor} is an interface which + allows to apply + the full expressiveness of ML-based theorem transformations + to code generation; motivating examples are shown below, see + \secref{sec:preproc}. + The result of the preprocessing step is a structured collection + of code equations. + + \item These code equations are \qn{translated} to a program + in an abstract intermediate language. Think of it as a kind + of \qt{Mini-Haskell} with four \qn{statements}: \isa{data} + (for datatypes), \isa{fun} (stemming from code equations), + also \isa{class} and \isa{inst} (for type classes). + + \item Finally, the abstract program is \qn{serialised} into concrete + source code of a target language. + + \end{itemize} + + \noindent From these steps, only the two last are carried out outside the logic; by + keeping this layer as thin as possible, the amount of code to trust is + kept to a minimum.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/document/ML.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/ML.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,255 @@ +% +\begin{isabellebody}% +\def\isabellecontext{ML}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ {\isachardoublequoteopen}ML{\isachardoublequoteclose}\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{ML system interfaces \label{sec:ml}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Since the code generator framework not only aims to provide + a nice Isar interface but also to form a base for + code-generation-based applications, here a short + description of the most important ML interfaces.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Executable theory content: \isa{Code}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This Pure module implements the core notions of + executable content of a theory.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Managing executable content% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\ + \indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\ + \indexdef{}{ML}{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\ + \indexdef{}{ML}{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\ + \indexdef{}{ML}{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\ + \indexdef{}{ML}{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline% +\verb| -> theory -> theory| \\ + \indexdef{}{ML}{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ + \indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ + \indexdef{}{ML}{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% +\verb| -> (string * sort) list * (string * typ list) list| \\ + \indexdef{}{ML}{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| + \end{mldecls} + + \begin{description} + + \item \verb|Code.add_eqn|~\isa{thm}~\isa{thy} adds function + theorem \isa{thm} to executable content. + + \item \verb|Code.del_eqn|~\isa{thm}~\isa{thy} removes function + theorem \isa{thm} from executable content, if present. + + \item \verb|Code.add_eqnl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds + suspended code equations \isa{lthms} for constant + \isa{const} to executable content. + + \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes + the preprocessor simpset. + + \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds + function transformer \isa{f} (named \isa{name}) to executable content; + \isa{f} is a transformer of the code equations belonging + to a certain function definition, depending on the + current theory context. Returning \isa{NONE} indicates that no + transformation took place; otherwise, the whole process will be iterated + with the new code equations. + + \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes + function transformer named \isa{name} from executable content. + + \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds + a datatype to executable content, with generation + set \isa{cs}. + + \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} + returns type constructor corresponding to + constructor \isa{const}; returns \isa{NONE} + if \isa{const} is no constructor. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Auxiliary% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ + \indexdef{}{ML}{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\ + \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\ + \end{mldecls} + + \begin{description} + + \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} + reads a constant as a concrete term expression \isa{s}. + + \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm} + extracts the constant and its type from a code equation \isa{thm}. + + \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm} + rewrites a code equation \isa{thm} with a simpset \isa{ss}; + only arguments and right hand side are rewritten, + not the head of the code equation. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Implementing code generator applications% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Implementing code generator applications on top + of the framework set out so far usually not only + involves using those primitive interfaces + but also storing code-dependent data and various + other things.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Data depending on the theory's executable content% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Due to incrementality of code generation, changes in the + theory's executable content have to be propagated in a + certain fashion. Additionally, such changes may occur + not only during theory extension but also during theory + merge, which is a little bit nasty from an implementation + point of view. The framework provides a solution + to this technical challenge by providing a functorial + data slot \verb|CodeDataFun|; on instantiation + of this functor, the following types and operations + are required: + + \medskip + \begin{tabular}{l} + \isa{type\ T} \\ + \isa{val\ empty{\isacharcolon}\ T} \\ + \isa{val\ purge{\isacharcolon}\ theory\ {\isasymrightarrow}\ string\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} + \end{tabular} + + \begin{description} + + \item \isa{T} the type of data to store. + + \item \isa{empty} initial (empty) data. + + \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; + \isa{consts} indicates the kind + of change: \verb|NONE| stands for a fundamental change + which invalidates any existing code, \isa{SOME\ consts} + hints that executable content for constants \isa{consts} + has changed. + + \end{description} + + \noindent An instance of \verb|CodeDataFun| provides the following + interface: + + \medskip + \begin{tabular}{l} + \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} + \end{tabular} + + \begin{description} + + \item \isa{get} retrieval of the current data. + + \item \isa{change} update of current data (cached!) + by giving a continuation. + + \item \isa{change{\isacharunderscore}yield} update with side result. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\bigskip + + \emph{Happy proving, happy hacking!}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/document/Program.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Program.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,1250 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Program}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Program\isanewline +\isakeyword{imports}\ Introduction\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Turning Theories into Programs \label{sec:program}% +} +\isamarkuptrue% +% +\isamarkupsubsection{The \isa{Isabelle{\isacharslash}HOL} default setup% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We have already seen how by default equations stemming from + \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} + statements are used for code generation. This default behaviour + can be changed, e.g. by providing different code equations. + All kinds of customisation shown in this section is \emph{safe} + in the sense that the user does not have to worry about + correctness -- all programs generatable that way are partially + correct.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Selecting code equations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Coming back to our introductory example, we + could provide an alternative code equations for \isa{dequeue} + explicitly:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}cases\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\ {\isacharparenleft}cases\ {\isachardoublequoteopen}rev\ xs{\isachardoublequoteclose}{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The annotation \isa{{\isacharbrackleft}code{\isacharbrackright}} is an \isa{Isar} + \isa{attribute} which states that the given theorems should be + considered as code equations for a \isa{fun} statement -- + the corresponding constant is determined syntactically. The resulting code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ +\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ +\hspace*{0pt}dequeue (AQueue xs []) =\\ +\hspace*{0pt} ~(if nulla xs then (Nothing,~AQueue [] [])\\ +\hspace*{0pt} ~~~else dequeue (AQueue [] (rev xs)));% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent You may note that the equality test \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} has been + replaced by the predicate \isa{null\ xs}. This is due to the default + setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). + + Changing the default constructor set of datatypes is also + possible. See \secref{sec:datatypes} for an example. + + As told in \secref{sec:concept}, code generation is based + on a structured collection of code theorems. + For explorative purpose, this collection + may be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% +\ dequeue% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent prints a table with \emph{all} code equations + for \isa{dequeue}, including + \emph{all} code equations those equations depend + on recursively. + + Similarly, the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command shows a graph + visualising dependencies between code equations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{\isa{class} and \isa{instantiation}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Concerning type classes and code generation, let us examine an example + from abstract algebra:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ semigroup\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{class}\isamarkupfalse% +\ monoid\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isasymotimes}\ n\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ n\ {\isacharplus}\ m\ {\isasymotimes}\ n{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ neutral{\isacharunderscore}nat\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharcolon}\isanewline +\ \ \isakeyword{fixes}\ n\ m\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ m{\isacharparenright}\ {\isasymotimes}\ q\ {\isacharequal}\ n\ {\isasymotimes}\ q\ {\isacharplus}\ m\ {\isasymotimes}\ q{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent We define the natural operation of the natural numbers + on monoids:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}pow\ {\isadigit{0}}\ a\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}pow\ {\isacharparenleft}Suc\ n{\isacharparenright}\ a\ {\isacharequal}\ a\ {\isasymotimes}\ pow\ n\ a{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This we use to define the discrete exponentiation function:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ bexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}bexp\ n\ {\isacharequal}\ pow\ n\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The corresponding code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}module Example where {\char123}\\ +\hspace*{0pt}\\ +\hspace*{0pt}\\ +\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}class Semigroup a where {\char123}\\ +\hspace*{0pt} ~mult ::~a -> a -> a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Semigroup a) => Monoid a where {\char123}\\ +\hspace*{0pt} ~neutral ::~a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\ +\hspace*{0pt}pow Zero{\char95}nat a = neutral;\\ +\hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\ +\hspace*{0pt}\\ +\hspace*{0pt}plus{\char95}nat ::~Nat -> Nat -> Nat;\\ +\hspace*{0pt}plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n);\\ +\hspace*{0pt}plus{\char95}nat Zero{\char95}nat n = n;\\ +\hspace*{0pt}\\ +\hspace*{0pt}neutral{\char95}nat ::~Nat;\\ +\hspace*{0pt}neutral{\char95}nat = Suc Zero{\char95}nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}mult{\char95}nat ::~Nat -> Nat -> Nat;\\ +\hspace*{0pt}mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat;\\ +\hspace*{0pt}mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Semigroup Nat where {\char123}\\ +\hspace*{0pt} ~mult = mult{\char95}nat;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Monoid Nat where {\char123}\\ +\hspace*{0pt} ~neutral = neutral{\char95}nat;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}bexp ::~Nat -> Nat;\\ +\hspace*{0pt}bexp n = pow n (Suc (Suc Zero{\char95}nat));\\ +\hspace*{0pt}\\ +\hspace*{0pt}{\char125}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This is a convenient place to show how explicit dictionary construction + manifests in generated code (here, the same example in \isa{SML}):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ +\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\ +\hspace*{0pt}fun semigroup{\char95}monoid (A{\char95}:'a monoid) = {\char35}Program{\char95}{\char95}semigroup{\char95}monoid A{\char95};\\ +\hspace*{0pt}fun neutral (A{\char95}:'a monoid) = {\char35}neutral A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun pow A{\char95}~Zero{\char95}nat a = neutral A{\char95}\\ +\hspace*{0pt} ~| pow A{\char95}~(Suc n) a = mult (semigroup{\char95}monoid A{\char95}) a (pow A{\char95}~n a);\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\ +\hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\ +\hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ +\hspace*{0pt}\\ +\hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoid{\char95}nat =\\ +\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}~:\\ +\hspace*{0pt} ~nat monoid;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Note the parameters with trailing underscore (\verb|A_|) + which are the dictionary parameters.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{The preprocessor \label{sec:preproc}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Before selected function theorems are turned into abstract + code, a chain of definitional transformation steps is carried + out: \emph{preprocessing}. In essence, the preprocessor + consists of two components: a \emph{simpset} and \emph{function transformers}. + + The \emph{simpset} allows to employ the full generality of the Isabelle + simplifier. Due to the interpretation of theorems + as code equations, rewrites are applied to the right + hand side and the arguments of the left hand side of an + equation, but never to the constant heading the left hand side. + An important special case are \emph{inline theorems} which may be + declared and undeclared using the + \emph{code inline} or \emph{code inline del} attribute respectively. + + Some common applications:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{itemize} +% +\begin{isamarkuptext}% +\item replacing non-executable constructs by executable ones:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\item eliminating superfluous constants:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\item replacing executable but inconvenient constructs:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\end{itemize} +% +\begin{isamarkuptext}% +\noindent \emph{Function transformers} provide a very general interface, + transforming a list of function theorems to another + list of function theorems, provided that neither the heading + constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} + pattern elimination implemented in + theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this + interface. + + \noindent The current setup of the preprocessor may be inspected using + the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. + \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} provides a convenient + mechanism to inspect the impact of a preprocessor setup + on code equations. + + \begin{warn} + The attribute \emph{code unfold} + associated with the \isa{SML\ code\ generator} also applies to + the \isa{generic\ code\ generator}: + \emph{code unfold} implies \emph{code inline}. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Datatypes \label{sec:datatypes}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Conceptually, any datatype is spanned by a set of + \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is exactly the set of \emph{all} type variables in + \isa{{\isasymtau}}. The HOL datatype package by default registers any new + datatype in the table of datatypes, which may be inspected using the + \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. + + In some cases, it is appropriate to alter or extend this table. As + an example, we will develop an alternative representation of the + queue example given in \secref{sec:intro}. The amortised + representation is convenient for generating code but exposes its + \qt{implementation} details, which may be cumbersome when proving + theorems about it. Therefore, here a simple, straightforward + representation of queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ {\isacharprime}a\ queue\ {\isacharequal}\ Queue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}Queue\ xs{\isacharparenright}\ {\isacharequal}\ Queue\ {\isacharparenleft}xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{fun}\isamarkupfalse% +\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ x{\isacharcomma}\ Queue\ xs{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This we can use directly for proving; for executing, + we provide an alternative characterisation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ AQueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}AQueue\ xs\ ys\ {\isacharequal}\ Queue\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% +\ AQueue% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Here we define a \qt{constructor} \isa{Program{\isachardot}AQueue} which + is defined in terms of \isa{Queue} and interprets its arguments + according to what the \emph{content} of an amortised queue is supposed + to be. Equipped with this, we are able to prove the following equations + for our primitive queue operations which \qt{implement} the simple + queues in an amortised fashion:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ empty{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ empty{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ enqueue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent For completeness, we provide a substitute for the + \isa{case} combinator on queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ aqueue{\isacharunderscore}case{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}aqueue{\isacharunderscore}case\ {\isacharequal}\ queue{\isacharunderscore}case{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ aqueue{\isacharunderscore}case\ {\isacharbrackleft}code{\isacharcomma}\ code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}queue{\isacharunderscore}case\ {\isacharequal}\ aqueue{\isacharunderscore}case{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ aqueue{\isacharunderscore}case{\isacharunderscore}def\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ case{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}aqueue{\isacharunderscore}case\ f\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ aqueue{\isacharunderscore}case{\isacharunderscore}def\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The resulting code looks as expected:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun foldl f a [] = a\\ +\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun null [] = true\\ +\hspace*{0pt} ~| null (x ::~xs) = false;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ +\hspace*{0pt} ~| dequeue (AQueue (xs,~[])) =\\ +\hspace*{0pt} ~~~(if null xs then (NONE,~AQueue ([],~[]))\\ +\hspace*{0pt} ~~~~~else dequeue (AQueue ([],~rev xs)));\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent From this example, it can be glimpsed that using own + constructor sets is a little delicate since it changes the set of + valid patterns for values of that type. Without going into much + detail, here some practical hints: + + \begin{itemize} + + \item When changing the constructor set for datatypes, take care + to provide an alternative for the \isa{case} combinator + (e.g.~by replacing it using the preprocessor). + + \item Values in the target language need not to be normalised -- + different values in the target language may represent the same + value in the logic. + + \item Usually, a good methodology to deal with the subtleties of + pattern matching is to see the type as an abstract type: provide + a set of operations which operate on the concrete representation + of the type, and derive further operations by combinations of + these primitive ones, without relying on a particular + representation. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Equality and wellsortedness% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Surely you have already noticed how equality is treated + by the code generator:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline +\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline +\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline +\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline +\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The membership test during preprocessing is rewritten, + resulting in \isa{op\ mem}, which itself + performs an explicit equality check.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun member A{\char95}~x [] = false\\ +\hspace*{0pt} ~| member A{\char95}~x (y ::~ys) = eqop A{\char95}~x y orelse member A{\char95}~x ys;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun collect{\char95}duplicates A{\char95}~xs ys [] = xs\\ +\hspace*{0pt} ~| collect{\char95}duplicates A{\char95}~xs ys (z ::~zs) =\\ +\hspace*{0pt} ~~~(if member A{\char95}~z xs\\ +\hspace*{0pt} ~~~~~then (if member A{\char95}~z ys then collect{\char95}duplicates A{\char95}~xs ys zs\\ +\hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\ +\hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Obviously, polymorphic equality is implemented the Haskell + way using a type class. How is this achieved? HOL introduces + an explicit class \isa{eq} with a corresponding operation + \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}. + The preprocessing framework does the rest by propagating the + \isa{eq} constraints through all dependent code equations. + For datatypes, instances of \isa{eq} are implicitly derived + when possible. For other types, you may instantiate \isa{eq} + manually like any other type class. + + Though this \isa{eq} class is designed to get rarely in + the way, a subtlety + enters the stage when definitions of overloaded constants + are dependent on operational equality. For example, let + us define a lexicographic ordering on tuples + (also see theory \hyperlink{theory.Product-ord}{\mbox{\isa{Product{\isacharunderscore}ord}}}):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}order{\isacharcomma}\ order{\isacharparenright}\ order\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}x\ {\isasymle}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isasymle}\ snd\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isacharless}\ snd\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}prod{\isacharunderscore}def\ intro{\isacharcolon}\ order{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ order{\isacharunderscore}prod\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then code generation will fail. Why? The definition + of \isa{op\ {\isasymle}} depends on equality on both arguments, + which are polymorphic and impose an additional \isa{eq} + class constraint, which the preprocessor does not propagate + (for technical reasons). + + The solution is to add \isa{eq} explicitly to the first sort arguments in the + code theorems:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ order{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then code generation succeeds:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\ +\hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a preorder = {\char123}Orderings{\char95}{\char95}ord{\char95}preorder :~'a ord{\char125};\\ +\hspace*{0pt}fun ord{\char95}preorder (A{\char95}:'a preorder) = {\char35}Orderings{\char95}{\char95}ord{\char95}preorder A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\ +\hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ +\hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ +\hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\ +\hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\ +\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ +\hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ +\hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\ +\hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +In some cases, the automatically derived code equations + for equality on a particular type may not be appropriate. + As example, watch the following datatype representing + monomorphic parametric types (where type constructors + are referred to by natural numbers):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Then code generation for SML would fail with a message + that the generated code contains illegal mutual dependencies: + the theorem \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}} already requires the + instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires + \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually + recursive \isa{instance} and \isa{function} definitions, + but the SML serialiser does not support this. + + In such cases, you have to provide your own equality equations + involving auxiliary constants. In our case, + \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun null [] = true\\ +\hspace*{0pt} ~| null (x ::~xs) = false;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eq{\char95}nat (Suc a) Zero{\char95}nat = false\\ +\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc a) = false\\ +\hspace*{0pt} ~| eq{\char95}nat (Suc nat) (Suc nat') = eq{\char95}nat nat nat'\\ +\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat Zero{\char95}nat = true;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype monotype = Mono of nat * monotype list;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun list{\char95}all2 p (x ::~xs) (y ::~ys) = p x y andalso list{\char95}all2 p xs ys\\ +\hspace*{0pt} ~| list{\char95}all2 p xs [] = null xs\\ +\hspace*{0pt} ~| list{\char95}all2 p [] ys = null ys;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1,~typargs1)) (Mono (tyco2,~typargs2)) =\\ +\hspace*{0pt} ~eq{\char95}nat tyco1 tyco2 andalso list{\char95}all2 eq{\char95}monotype typargs1 typargs2;\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsubsection{Explicit partiality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Partiality usually enters the game by partial patterns, as + in the following example, again for amortised queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\isanewline +\ \ \ \ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent In the corresponding code, there is no equation + for the pattern \isa{Program{\isachardot}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\ +\hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\ +\hspace*{0pt} ~let {\char123}\\ +\hspace*{0pt} ~~~(y :~ys) = rev xs;\\ +\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\ +\hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent In some cases it is desirable to have this + pseudo-\qt{partiality} more explicitly, e.g.~as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{axiomatization}\isamarkupfalse% +\ empty{\isacharunderscore}queue\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isacharbar}\ {\isacharunderscore}\ {\isasymRightarrow}\ empty{\isacharunderscore}queue{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ empty{\isacharunderscore}queue\isanewline +\ \ \ \ \ else\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +Observe that on the right hand side of the definition of \isa{strict{\isacharunderscore}dequeue{\isacharprime}} the constant \isa{empty{\isacharunderscore}queue} occurs + which is unspecified. + + Normally, if constants without any code equations occur in a + program, the code generator complains (since in most cases this is + not what the user expects). But such constants can also be thought + of as function definitions with no equations which always fail, + since there is never a successful pattern match on the left hand + side. In order to categorise a constant into that category + explicitly, use \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}abort}\isamarkupfalse% +\ empty{\isacharunderscore}queue% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then the code generator will just insert an error or + exception at the appropriate position:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\ +\hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\ +\hspace*{0pt}\\ +\hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\ +\hspace*{0pt}strict{\char95}dequeue' (AQueue xs (y :~ys)) = (y,~AQueue xs ys);\\ +\hspace*{0pt}strict{\char95}dequeue' (AQueue xs []) =\\ +\hspace*{0pt} ~(if nulla xs then empty{\char95}queue\\ +\hspace*{0pt} ~~~else strict{\char95}dequeue' (AQueue [] (rev xs)));% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This feature however is rarely needed in practice. + Note also that the \isa{HOL} default setup already declares + \isa{undefined} as \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}, which is most + likely to be used in such situations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\ \end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/Codegen.hs --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/Codegen.hs Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,23 @@ +module Codegen where { + +import qualified Nat; + +class Null a where { + nulla :: a; +}; + +heada :: forall a. (Codegen.Null a) => [a] -> a; +heada (x : xs) = x; +heada [] = Codegen.nulla; + +null_option :: forall a. Maybe a; +null_option = Nothing; + +instance Codegen.Null (Maybe a) where { + nulla = Codegen.null_option; +}; + +dummy :: Maybe Nat.Nat; +dummy = Codegen.heada [Just (Nat.Suc Nat.Zero_nat), Nothing]; + +} diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/Example.hs --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/Example.hs Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,33 @@ +{-# OPTIONS_GHC -fglasgow-exts #-} + +module Example where { + + +foldla :: forall a b. (a -> b -> a) -> a -> [b] -> a; +foldla f a [] = a; +foldla f a (x : xs) = foldla f (f a x) xs; + +rev :: forall a. [a] -> [a]; +rev xs = foldla (\ xsa x -> x : xsa) [] xs; + +list_case :: forall t a. t -> (a -> [a] -> t) -> [a] -> t; +list_case f1 f2 (a : list) = f2 a list; +list_case f1 f2 [] = f1; + +data Queue a = AQueue [a] [a]; + +empty :: forall a. Queue a; +empty = AQueue [] []; + +dequeue :: forall a. Queue a -> (Maybe a, Queue a); +dequeue (AQueue [] []) = (Nothing, AQueue [] []); +dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys); +dequeue (AQueue (v : va) []) = + let { + (y : ys) = rev (v : va); + } in (Just y, AQueue [] ys); + +enqueue :: forall a. a -> Queue a -> Queue a; +enqueue x (AQueue xs ys) = AQueue (x : xs) ys; + +} diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/arbitrary.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/arbitrary.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,9 @@ +structure Codegen = +struct + +val arbitrary_option : 'a option = NONE; + +fun dummy_option [] = arbitrary_option + | dummy_option (x :: xs) = SOME x; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/bool_infix.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/bool_infix.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,19 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun in_interval (k, l) n = + Nat.less_eq_nat k n andalso Nat.less_eq_nat n l; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/bool_literal.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/bool_literal.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,31 @@ +structure HOL = +struct + +datatype boola = False | True; + +fun anda x True = x + | anda x False = False + | anda True x = x + | anda False x = False; + +end; (*struct HOL*) + +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = HOL.False +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = HOL.True; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun in_interval (k, l) n = + HOL.anda (Nat.less_eq_nat k n) (Nat.less_eq_nat n l); + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/bool_mlbool.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/bool_mlbool.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,19 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun in_interval (k, l) n = + (Nat.less_eq_nat k n) andalso (Nat.less_eq_nat n l); + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/class.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/class.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,24 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +end; (*struct Nat*) + +structure Codegen = +struct + +type 'a null = {null : 'a}; +fun null (A_:'a null) = #null A_; + +fun head A_ (x :: xs) = x + | head A_ [] = null A_; + +val null_option : 'a option = NONE; + +fun null_optiona () = {null = null_option} : ('a option) null; + +val dummy : Nat.nat option = + head (null_optiona ()) [SOME (Nat.Suc Nat.Zero_nat), NONE]; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/class.ocaml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/class.ocaml Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,24 @@ +module Nat = +struct + +type nat = Suc of nat | Zero_nat;; + +end;; (*struct Nat*) + +module Codegen = +struct + +type 'a null = {null : 'a};; +let null _A = _A.null;; + +let rec head _A = function x :: xs -> x + | [] -> null _A;; + +let rec null_option = None;; + +let null_optiona () = ({null = null_option} : ('a option) null);; + +let rec dummy + = head (null_optiona ()) [Some (Nat.Suc Nat.Zero_nat); None];; + +end;; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/collect_duplicates.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/collect_duplicates.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,30 @@ +structure HOL = +struct + +type 'a eq = {eq : 'a -> 'a -> bool}; +fun eq (A_:'a eq) = #eq A_; + +fun eqop A_ a = eq A_ a; + +end; (*struct HOL*) + +structure List = +struct + +fun member A_ x (y :: ys) = + (if HOL.eqop A_ y x then true else member A_ x ys) + | member A_ x [] = false; + +end; (*struct List*) + +structure Codegen = +struct + +fun collect_duplicates A_ xs ys (z :: zs) = + (if List.member A_ z xs + then (if List.member A_ z ys then collect_duplicates A_ xs ys zs + else collect_duplicates A_ xs (z :: ys) zs) + else collect_duplicates A_ (z :: xs) (z :: ys) zs) + | collect_duplicates A_ xs ys [] = xs; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/dirty_set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/dirty_set.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,102 @@ +structure ROOT = +struct + +structure Nat = +struct + +datatype nat = Zero_nat | Suc of nat; + +end; (*struct Nat*) + +structure Integer = +struct + +datatype bit = B0 | B1; + +datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; + +fun pred (Bit (k, B0)) = Bit (pred k, B1) + | pred (Bit (k, B1)) = Bit (k, B0) + | pred Min = Bit (Min, B0) + | pred Pls = Min; + +fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) + | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) + | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) + | uminus_int Min = Bit (Pls, B1) + | uminus_int Pls = Pls; + +fun succ (Bit (k, B0)) = Bit (k, B1) + | succ (Bit (k, B1)) = Bit (succ k, B0) + | succ Min = Pls + | succ Pls = Bit (Pls, B1); + +fun plus_int (Number_of_int v) (Number_of_int w) = + Number_of_int (plus_int v w) + | plus_int k Min = pred k + | plus_int k Pls = k + | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) + | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) + | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) + | plus_int Min k = pred k + | plus_int Pls k = k; + +fun minus_int (Number_of_int v) (Number_of_int w) = + Number_of_int (plus_int v (uminus_int w)) + | minus_int z w = plus_int z (uminus_int w); + +fun less_eq_int (Number_of_int k) (Number_of_int l) = less_eq_int k l + | less_eq_int (Bit (k1, B1)) (Bit (k2, B0)) = less_int k1 k2 + | less_eq_int (Bit (k1, v)) (Bit (k2, B1)) = less_eq_int k1 k2 + | less_eq_int (Bit (k1, B0)) (Bit (k2, v)) = less_eq_int k1 k2 + | less_eq_int (Bit (k, v)) Min = less_eq_int k Min + | less_eq_int (Bit (k, B1)) Pls = less_int k Pls + | less_eq_int (Bit (k, B0)) Pls = less_eq_int k Pls + | less_eq_int Min (Bit (k, B1)) = less_eq_int Min k + | less_eq_int Min (Bit (k, B0)) = less_int Min k + | less_eq_int Min Min = true + | less_eq_int Min Pls = true + | less_eq_int Pls (Bit (k, v)) = less_eq_int Pls k + | less_eq_int Pls Min = false + | less_eq_int Pls Pls = true +and less_int (Number_of_int k) (Number_of_int l) = less_int k l + | less_int (Bit (k1, B0)) (Bit (k2, B1)) = less_eq_int k1 k2 + | less_int (Bit (k1, B1)) (Bit (k2, v)) = less_int k1 k2 + | less_int (Bit (k1, v)) (Bit (k2, B0)) = less_int k1 k2 + | less_int (Bit (k, B1)) Min = less_int k Min + | less_int (Bit (k, B0)) Min = less_eq_int k Min + | less_int (Bit (k, v)) Pls = less_int k Pls + | less_int Min (Bit (k, v)) = less_int Min k + | less_int Min Min = false + | less_int Min Pls = true + | less_int Pls (Bit (k, B1)) = less_eq_int Pls k + | less_int Pls (Bit (k, B0)) = less_int Pls k + | less_int Pls Min = false + | less_int Pls Pls = false; + +fun nat_aux n i = + (if less_eq_int i (Number_of_int Pls) then n + else nat_aux (Nat.Suc n) + (minus_int i (Number_of_int (Bit (Pls, B1))))); + +fun nat i = nat_aux Nat.Zero_nat i; + +end; (*struct Integer*) + +structure Codegen = +struct + +val dummy_set : (Nat.nat -> Nat.nat) list = Nat.Suc :: []; + +val foobar_set : Nat.nat list = + Nat.Zero_nat :: + (Nat.Suc Nat.Zero_nat :: + (Integer.nat + (Integer.Number_of_int + (Integer.Bit + (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) + :: [])); + +end; (*struct Codegen*) + +end; (*struct ROOT*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/example.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/example.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,27 @@ +structure Example = +struct + +fun foldl f a [] = a + | foldl f a (x :: xs) = foldl f (f a x) xs; + +fun rev xs = foldl (fn xsa => fn x => x :: xsa) [] xs; + +fun list_case f1 f2 (a :: lista) = f2 a lista + | list_case f1 f2 [] = f1; + +datatype 'a queue = AQueue of 'a list * 'a list; + +val empty : 'a queue = AQueue ([], []) + +fun dequeue (AQueue ([], [])) = (NONE, AQueue ([], [])) + | dequeue (AQueue (xs, y :: ys)) = (SOME y, AQueue (xs, ys)) + | dequeue (AQueue (v :: va, [])) = + let + val y :: ys = rev (v :: va); + in + (SOME y, AQueue ([], ys)) + end; + +fun enqueue x (AQueue (xs, ys)) = AQueue (x :: xs, ys); + +end; (*struct Example*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/fac.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/fac.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,22 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +val one_nat : nat = Suc Zero_nat; + +fun plus_nat (Suc m) n = plus_nat m (Suc n) + | plus_nat Zero_nat n = n; + +fun times_nat (Suc m) n = plus_nat n (times_nat m n) + | times_nat Zero_nat n = Zero_nat; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun fac (Nat.Suc n) = Nat.times_nat (Nat.Suc n) (fac n) + | fac Nat.Zero_nat = Nat.one_nat; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/integers.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/integers.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,59 @@ +structure ROOT = +struct + +structure Integer = +struct + +datatype bit = B0 | B1; + +datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; + +fun pred (Bit (k, B0)) = Bit (pred k, B1) + | pred (Bit (k, B1)) = Bit (k, B0) + | pred Min = Bit (Min, B0) + | pred Pls = Min; + +fun succ (Bit (k, B0)) = Bit (k, B1) + | succ (Bit (k, B1)) = Bit (succ k, B0) + | succ Min = Pls + | succ Pls = Bit (Pls, B1); + +fun plus_int (Number_of_int v) (Number_of_int w) = + Number_of_int (plus_int v w) + | plus_int k Min = pred k + | plus_int k Pls = k + | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) + | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) + | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) + | plus_int Min k = pred k + | plus_int Pls k = k; + +fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) + | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) + | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) + | uminus_int Min = Bit (Pls, B1) + | uminus_int Pls = Pls; + +fun times_int (Number_of_int v) (Number_of_int w) = + Number_of_int (times_int v w) + | times_int (Bit (k, B0)) l = Bit (times_int k l, B0) + | times_int (Bit (k, B1)) l = plus_int (Bit (times_int k l, B0)) l + | times_int Min k = uminus_int k + | times_int Pls w = Pls; + +end; (*struct Integer*) + +structure Codegen = +struct + +fun double_inc k = + Integer.plus_int + (Integer.times_int + (Integer.Number_of_int + (Integer.Bit (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) + k) + (Integer.Number_of_int (Integer.Bit (Integer.Pls, Integer.B1))); + +end; (*struct Codegen*) + +end; (*struct ROOT*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/lexicographic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/lexicographic.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,19 @@ +structure HOL = +struct + +type 'a eq = {eq : 'a -> 'a -> bool}; +fun eq (A_:'a eq) = #eq A_; + +type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; +fun less_eq (A_:'a ord) = #less_eq A_; +fun less (A_:'a ord) = #less A_; + +end; (*struct HOL*) + +structure Codegen = +struct + +fun less_eq (A1_, A2_) B_ (x1, y1) (x2, y2) = + HOL.less A2_ x1 x2 orelse HOL.eq A1_ x1 x2 andalso HOL.less_eq B_ y1 y2; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/lookup.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/lookup.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,13 @@ +structure ROOT = +struct + +structure Codegen = +struct + +fun lookup ((k, v) :: xs) l = + (if ((k : string) = l) then SOME v else lookup xs l) + | lookup [] l = NONE; + +end; (*struct Codegen*) + +end; (*struct ROOT*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/monotype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/monotype.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,34 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun eq_nat (Suc a) Zero_nat = false + | eq_nat Zero_nat (Suc a) = false + | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' + | eq_nat Zero_nat Zero_nat = true; + +end; (*struct Nat*) + +structure List = +struct + +fun null (x :: xs) = false + | null [] = true; + +fun list_all2 p (x :: xs) (y :: ys) = p x y andalso list_all2 p xs ys + | list_all2 p xs [] = null xs + | list_all2 p [] ys = null ys; + +end; (*struct List*) + +structure Codegen = +struct + +datatype monotype = Mono of Nat.nat * monotype list; + +fun eq_monotype (Mono (tyco1, typargs1)) (Mono (tyco2, typargs2)) = + Nat.eq_nat tyco1 tyco2 andalso + List.list_all2 eq_monotype typargs1 typargs2; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/nat_binary.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/nat_binary.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,17 @@ +structure Nat = +struct + +datatype nat = Dig1 of nat | Dig0 of nat | One_nat | Zero_nat; + +fun plus_nat (Dig1 m) (Dig1 n) = Dig0 (plus_nat (plus_nat m n) One_nat) + | plus_nat (Dig1 m) (Dig0 n) = Dig1 (plus_nat m n) + | plus_nat (Dig0 m) (Dig1 n) = Dig1 (plus_nat m n) + | plus_nat (Dig0 m) (Dig0 n) = Dig0 (plus_nat m n) + | plus_nat (Dig1 m) One_nat = Dig0 (plus_nat m One_nat) + | plus_nat One_nat (Dig1 n) = Dig0 (plus_nat n One_nat) + | plus_nat (Dig0 m) One_nat = Dig1 m + | plus_nat One_nat (Dig0 n) = Dig1 n + | plus_nat m Zero_nat = m + | plus_nat Zero_nat n = n; + +end; (*struct Nat*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/pick1.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/pick1.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,44 @@ +structure HOL = +struct + +fun leta s f = f s; + +end; (*struct HOL*) + +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +fun minus_nat (Suc m) (Suc n) = minus_nat m n + | minus_nat Zero_nat n = Zero_nat + | minus_nat m Zero_nat = m; + +end; (*struct Nat*) + +structure Product_Type = +struct + +fun split f (a, b) = f a b; + +end; (*struct Product_Type*) + +structure Codegen = +struct + +fun pick ((k, v) :: xs) n = + (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) + | pick (x :: xs) n = + let + val a = x; + val (k, v) = a; + in + (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) + end; + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/Thy/examples/tree.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/tree.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,95 @@ +structure HOL = +struct + +type 'a eq = {eq : 'a -> 'a -> bool}; +fun eq (A_:'a eq) = #eq A_; + +type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; +fun less_eq (A_:'a ord) = #less_eq A_; +fun less (A_:'a ord) = #less A_; + +fun eqop A_ a = eq A_ a; + +end; (*struct HOL*) + +structure Orderings = +struct + +type 'a preorder = {Orderings__ord_preorder : 'a HOL.ord}; +fun ord_preorder (A_:'a preorder) = #Orderings__ord_preorder A_; + +type 'a order = {Orderings__preorder_order : 'a preorder}; +fun preorder_order (A_:'a order) = #Orderings__preorder_order A_; + +type 'a linorder = {Orderings__order_linorder : 'a order}; +fun order_linorder (A_:'a linorder) = #Orderings__order_linorder A_; + +end; (*struct Orderings*) + +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun eq_nat (Suc a) Zero_nat = false + | eq_nat Zero_nat (Suc a) = false + | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' + | eq_nat Zero_nat Zero_nat = true; + +val eq_nata = {eq = eq_nat} : nat HOL.eq; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +val ord_nat = {less_eq = less_eq_nat, less = less_nat} : nat HOL.ord; + +val preorder_nat = {Orderings__ord_preorder = ord_nat} : + nat Orderings.preorder; + +val order_nat = {Orderings__preorder_order = preorder_nat} : + nat Orderings.order; + +val linorder_nat = {Orderings__order_linorder = order_nat} : + nat Orderings.linorder; + +end; (*struct Nat*) + +structure Codegen = +struct + +datatype ('a, 'b) searchtree = + Branch of ('a, 'b) searchtree * 'a * ('a, 'b) searchtree | + Leaf of 'a * 'b; + +fun update (A1_, A2_) (it, entry) (Leaf (key, vala)) = + (if HOL.eqop A1_ it key then Leaf (key, entry) + else (if HOL.less_eq + ((Orderings.ord_preorder o Orderings.preorder_order o + Orderings.order_linorder) + A2_) + it key + then Branch (Leaf (it, entry), it, Leaf (key, vala)) + else Branch (Leaf (key, vala), it, Leaf (it, entry)))) + | update (A1_, A2_) (it, entry) (Branch (t1, key, t2)) = + (if HOL.less_eq + ((Orderings.ord_preorder o Orderings.preorder_order o + Orderings.order_linorder) + A2_) + it key + then Branch (update (A1_, A2_) (it, entry) t1, key, t2) + else Branch (t1, key, update (A1_, A2_) (it, entry) t2)); + +val example : (Nat.nat, (Nat.nat list)) searchtree = + update (Nat.eq_nata, Nat.linorder_nat) + (Nat.Suc (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))), + [Nat.Suc (Nat.Suc Nat.Zero_nat), Nat.Suc (Nat.Suc Nat.Zero_nat)]) + (update (Nat.eq_nata, Nat.linorder_nat) + (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat)), + [Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))]) + (update (Nat.eq_nata, Nat.linorder_nat) + (Nat.Suc (Nat.Suc Nat.Zero_nat), [Nat.Suc (Nat.Suc Nat.Zero_nat)]) + (Leaf (Nat.Suc Nat.Zero_nat, [])))); + +end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/codegen.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/codegen.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,53 @@ + +\documentclass[12pt,a4paper,fleqn]{report} +\usepackage{latexsym,graphicx} +\usepackage[refpage]{nomencl} +\usepackage{../iman,../extra,../isar,../proof} +\usepackage{../isabelle,../isabellesym} +\usepackage{style} +\usepackage{pgf} +\usepackage{pgflibraryshapes} +\usepackage{tikz} +\usepackage{../pdfsetup} + +\hyphenation{Isabelle} +\hyphenation{Isar} +\isadroptag{theory} + +\title{\includegraphics[scale=0.5]{isabelle_isar} + \\[4ex] Code generation from Isabelle/HOL theories} +\author{\emph{Florian Haftmann}} + +\begin{document} + +\maketitle + +\begin{abstract} + This tutorial gives an introduction to a generic code generator framework in Isabelle + for generating executable code in functional programming languages from logical + specifications in Isabelle/HOL. +\end{abstract} + +\thispagestyle{empty}\clearpage + +\pagenumbering{roman} +\clearfirst + +\input{Thy/document/Introduction.tex} +\input{Thy/document/Program.tex} +\input{Thy/document/Adaption.tex} +\input{Thy/document/Further.tex} +\input{Thy/document/ML.tex} + +\begingroup +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{../manual} +\endgroup + +\end{document} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/codegen_process.pdf Binary file doc-src/Codegen/codegen_process.pdf has changed diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/codegen_process.ps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/codegen_process.ps Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,586 @@ +%!PS-Adobe-2.0 +%%Creator: dot version 2.2 (Mon Sep 12 23:33:36 UTC 2005) +%%For: (haftmann) Florian Haftmann +%%Title: _anonymous_0 +%%Pages: (atend) +%%BoundingBox: 35 35 451 291 +%%EndComments +save +%%BeginProlog +/DotDict 200 dict def +DotDict begin + +/setupLatin1 { +mark +/EncodingVector 256 array def + EncodingVector 0 + +ISOLatin1Encoding 0 255 getinterval putinterval + +EncodingVector + dup 306 /AE + dup 301 /Aacute + dup 302 /Acircumflex + dup 304 /Adieresis + dup 300 /Agrave + dup 305 /Aring + dup 303 /Atilde + dup 307 /Ccedilla + dup 311 /Eacute + dup 312 /Ecircumflex + dup 313 /Edieresis + dup 310 /Egrave + dup 315 /Iacute + dup 316 /Icircumflex + dup 317 /Idieresis + dup 314 /Igrave + dup 334 /Udieresis + dup 335 /Yacute + dup 376 /thorn + dup 337 /germandbls + dup 341 /aacute + dup 342 /acircumflex + dup 344 /adieresis + dup 346 /ae + dup 340 /agrave + dup 345 /aring + dup 347 /ccedilla + dup 351 /eacute + dup 352 /ecircumflex + dup 353 /edieresis + dup 350 /egrave + dup 355 /iacute + dup 356 /icircumflex + dup 357 /idieresis + dup 354 /igrave + dup 360 /dcroat + dup 361 /ntilde + dup 363 /oacute + dup 364 /ocircumflex + dup 366 /odieresis + dup 362 /ograve + dup 365 /otilde + dup 370 /oslash + dup 372 /uacute + dup 373 /ucircumflex + dup 374 /udieresis + dup 371 /ugrave + dup 375 /yacute + dup 377 /ydieresis + +% Set up ISO Latin 1 character encoding +/starnetISO { + dup dup findfont dup length dict begin + { 1 index /FID ne { def }{ pop pop } ifelse + } forall + /Encoding EncodingVector def + currentdict end definefont +} def +/Times-Roman starnetISO def +/Times-Italic starnetISO def +/Times-Bold starnetISO def +/Times-BoldItalic starnetISO def +/Helvetica starnetISO def +/Helvetica-Oblique starnetISO def +/Helvetica-Bold starnetISO def +/Helvetica-BoldOblique starnetISO def +/Courier starnetISO def +/Courier-Oblique starnetISO def +/Courier-Bold starnetISO def +/Courier-BoldOblique starnetISO def +cleartomark +} bind def + +%%BeginResource: procset graphviz 0 0 +/coord-font-family /Times-Roman def +/default-font-family /Times-Roman def +/coordfont coord-font-family findfont 8 scalefont def + +/InvScaleFactor 1.0 def +/set_scale { + dup 1 exch div /InvScaleFactor exch def + dup scale +} bind def + +% styles +/solid { [] 0 setdash } bind def +/dashed { [9 InvScaleFactor mul dup ] 0 setdash } bind def +/dotted { [1 InvScaleFactor mul 6 InvScaleFactor mul] 0 setdash } bind def +/invis {/fill {newpath} def /stroke {newpath} def /show {pop newpath} def} bind def +/bold { 2 setlinewidth } bind def +/filled { } bind def +/unfilled { } bind def +/rounded { } bind def +/diagonals { } bind def + +% hooks for setting color +/nodecolor { sethsbcolor } bind def +/edgecolor { sethsbcolor } bind def +/graphcolor { sethsbcolor } bind def +/nopcolor {pop pop pop} bind def + +/beginpage { % i j npages + /npages exch def + /j exch def + /i exch def + /str 10 string def + npages 1 gt { + gsave + coordfont setfont + 0 0 moveto + (\() show i str cvs show (,) show j str cvs show (\)) show + grestore + } if +} bind def + +/set_font { + findfont exch + scalefont setfont +} def + +% draw aligned label in bounding box aligned to current point +/alignedtext { % width adj text + /text exch def + /adj exch def + /width exch def + gsave + width 0 gt { + text stringwidth pop adj mul 0 rmoveto + } if + [] 0 setdash + text show + grestore +} def + +/boxprim { % xcorner ycorner xsize ysize + 4 2 roll + moveto + 2 copy + exch 0 rlineto + 0 exch rlineto + pop neg 0 rlineto + closepath +} bind def + +/ellipse_path { + /ry exch def + /rx exch def + /y exch def + /x exch def + matrix currentmatrix + newpath + x y translate + rx ry scale + 0 0 1 0 360 arc + setmatrix +} bind def + +/endpage { showpage } bind def +/showpage { } def + +/layercolorseq + [ % layer color sequence - darkest to lightest + [0 0 0] + [.2 .8 .8] + [.4 .8 .8] + [.6 .8 .8] + [.8 .8 .8] + ] +def + +/layerlen layercolorseq length def + +/setlayer {/maxlayer exch def /curlayer exch def + layercolorseq curlayer 1 sub layerlen mod get + aload pop sethsbcolor + /nodecolor {nopcolor} def + /edgecolor {nopcolor} def + /graphcolor {nopcolor} def +} bind def + +/onlayer { curlayer ne {invis} if } def + +/onlayers { + /myupper exch def + /mylower exch def + curlayer mylower lt + curlayer myupper gt + or + {invis} if +} def + +/curlayer 0 def + +%%EndResource +%%EndProlog +%%BeginSetup +14 default-font-family set_font +1 setmiterlimit +% /arrowlength 10 def +% /arrowwidth 5 def + +% make sure pdfmark is harmless for PS-interpreters other than Distiller +/pdfmark where {pop} {userdict /pdfmark /cleartomark load put} ifelse +% make '<<' and '>>' safe on PS Level 1 devices +/languagelevel where {pop languagelevel}{1} ifelse +2 lt { + userdict (<<) cvn ([) cvn load put + userdict (>>) cvn ([) cvn load put +} if + +%%EndSetup +%%Page: 1 1 +%%PageBoundingBox: 36 36 451 291 +%%PageOrientation: Portrait +gsave +35 35 416 256 boxprim clip newpath +36 36 translate +0 0 1 beginpage +0 0 translate 0 rotate +[ /CropBox [36 36 451 291] /PAGES pdfmark +0.000 0.000 0.000 graphcolor +14.00 /Times-Roman set_font + +% theory +gsave 10 dict begin +newpath 93 254 moveto +1 254 lineto +1 214 lineto +93 214 lineto +closepath +stroke +gsave 10 dict begin +8 237 moveto +(Isabelle/HOL) +[4.56 5.52 6.24 6.96 6.24 3.84 3.84 6.24 3.84 10.08 10.08 8.64] +xshow +16 221 moveto +(Isar theory) +[4.56 5.52 6.24 4.56 3.6 4.08 6.96 6.24 6.96 4.8 6.96] +xshow +end grestore +end grestore + +% selection +gsave 10 dict begin +183 234 38 18 ellipse_path +stroke +gsave 10 dict begin +158 229 moveto +(selection) +[5.52 6.24 3.84 6.24 6.24 3.84 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% theory -> selection +newpath 94 234 moveto +107 234 121 234 135 234 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 135 238 moveto +145 234 lineto +135 231 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 135 238 moveto +145 234 lineto +135 231 lineto +closepath +stroke +end grestore + +% sml +gsave 10 dict begin +newpath 74 144 moveto +20 144 lineto +20 108 lineto +74 108 lineto +closepath +stroke +gsave 10 dict begin +32 121 moveto +(SML) +[7.68 12.48 8.64] +xshow +end grestore +end grestore + +% other +gsave 10 dict begin +gsave 10 dict begin +41 67 moveto +(...) +[3.6 3.6 3.6] +xshow +end grestore +end grestore + +% haskell +gsave 10 dict begin +newpath 77 36 moveto +17 36 lineto +17 0 lineto +77 0 lineto +closepath +stroke +gsave 10 dict begin +25 13 moveto +(Haskell) +[10.08 6.24 5.52 6.72 6.24 3.84 3.84] +xshow +end grestore +end grestore + +% preprocessing +gsave 10 dict begin +183 180 52 18 ellipse_path +stroke +gsave 10 dict begin +143 175 moveto +(preprocessing) +[6.96 4.56 6.24 6.96 4.56 6.96 6.24 6.24 5.52 5.52 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% selection -> preprocessing +newpath 183 216 moveto +183 213 183 211 183 208 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 187 208 moveto +183 198 lineto +180 208 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 187 208 moveto +183 198 lineto +180 208 lineto +closepath +stroke +end grestore + +% def_eqn +gsave 10 dict begin +newpath 403 198 moveto +283 198 lineto +283 162 lineto +403 162 lineto +closepath +stroke +gsave 10 dict begin +291 175 moveto +(defining equations) +[6.96 6.24 4.8 3.84 6.96 3.84 6.96 6.96 3.6 6.24 6.72 6.96 6.24 3.84 3.84 6.96 6.96 5.52] +xshow +end grestore +end grestore + +% preprocessing -> def_eqn +newpath 236 180 moveto +248 180 260 180 273 180 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 273 184 moveto +283 180 lineto +273 177 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 273 184 moveto +283 180 lineto +273 177 lineto +closepath +stroke +end grestore + +% serialization +gsave 10 dict begin +183 72 47 18 ellipse_path +stroke +gsave 10 dict begin +148 67 moveto +(serialization) +[5.52 6.24 4.8 3.84 6.24 3.84 3.84 6.24 6.24 3.84 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% serialization -> sml +newpath 150 85 moveto +129 93 104 103 83 111 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 82 108 moveto +74 115 lineto +85 114 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 82 108 moveto +74 115 lineto +85 114 lineto +closepath +stroke +end grestore + +% serialization -> other +gsave 10 dict begin +dotted +newpath 135 72 moveto +119 72 100 72 84 72 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 84 69 moveto +74 72 lineto +84 76 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 84 69 moveto +74 72 lineto +84 76 lineto +closepath +stroke +end grestore +end grestore + +% serialization -> haskell +newpath 150 59 moveto +131 51 107 42 86 34 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 88 31 moveto +77 30 lineto +85 37 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 88 31 moveto +77 30 lineto +85 37 lineto +closepath +stroke +end grestore + +% translation +gsave 10 dict begin +343 126 43 18 ellipse_path +stroke +gsave 10 dict begin +313 121 moveto +(translation) +[3.84 4.56 6.24 6.96 5.52 3.84 6.24 3.84 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% def_eqn -> translation +newpath 343 162 moveto +343 159 343 157 343 154 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 347 154 moveto +343 144 lineto +340 154 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 347 154 moveto +343 144 lineto +340 154 lineto +closepath +stroke +end grestore + +% iml +gsave 10 dict begin +newpath 413 90 moveto +273 90 lineto +273 54 lineto +413 54 lineto +closepath +stroke +gsave 10 dict begin +280 67 moveto +(intermediate language) +[3.84 6.96 3.84 6.24 4.8 10.8 6.24 6.96 3.84 6.24 3.84 6.24 3.6 3.84 6.24 6.96 6.96 6.96 6.24 6.72 6.24] +xshow +end grestore +end grestore + +% translation -> iml +newpath 343 108 moveto +343 105 343 103 343 100 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 347 100 moveto +343 90 lineto +340 100 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 347 100 moveto +343 90 lineto +340 100 lineto +closepath +stroke +end grestore + +% iml -> serialization +newpath 272 72 moveto +262 72 251 72 241 72 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 241 69 moveto +231 72 lineto +241 76 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 241 69 moveto +231 72 lineto +241 76 lineto +closepath +stroke +end grestore +endpage +showpage +grestore +%%PageTrailer +%%EndPage: 1 +%%Trailer +%%Pages: 1 +end +restore +%%EOF diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Codegen/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/style.sty Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,56 @@ + +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} + +%% logical markup +\newcommand{\strong}[1]{{\bfseries {#1}}} +\newcommand{\qn}[1]{\emph{#1}} + +%% typographic conventions +\newcommand{\qt}[1]{``{#1}''} + +%% verbatim text +\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} + +%% quoted segments +\makeatletter +\isakeeptag{quote} +\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}} +\renewcommand{\isatagquote}{\begin{quotesegment}} +\renewcommand{\endisatagquote}{\end{quotesegment}} +\makeatother + +\isakeeptag{quotett} +\renewcommand{\isatagquotett}{\begin{quotesegment}\isabellestyle{tt}\isastyle} +\renewcommand{\endisatagquotett}{\end{quotesegment}\isabellestyle{it}\isastyle} + +%% a trick +\newcommand{\isasymSML}{SML} + +%% presentation +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +\pagestyle{headings} +\binperiod +\underscoreoff + +\renewcommand{\isadigit}[1]{\isamath{#1}} + +%% ml reference +\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} + +\isakeeptag{mlref} +\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} +\renewcommand{\endisatagmlref}{\endgroup} + +\isabellestyle{it} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "implementation" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Dirs --- a/doc-src/Dirs Mon Mar 02 16:58:39 2009 +0100 +++ b/doc-src/Dirs Tue Mar 03 11:00:51 2009 +0100 @@ -1,1 +1,1 @@ -Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar IsarAdvanced/Classes IsarAdvanced/Codegen IsarAdvanced/Functions +Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar Classes Codegen Functions diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/IsaMakefile Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,33 @@ + +## targets + +default: Thy +images: +test: Thy + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document + + +## Thy + +THY = $(LOG)/HOL-Thy.gz + +Thy: $(THY) + +$(THY): Thy/ROOT.ML Thy/Functions.thy + @$(USEDIR) HOL Thy + + +## clean + +clean: + @rm -f $(THY) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Makefile Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,38 @@ +# +# $Id$ +# + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = functions + +FILES = $(NAME).tex Thy/document/Functions.tex intro.tex conclusion.tex \ + style.sty ../iman.sty ../extra.sty ../isar.sty \ + ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ + ../manual.bib ../proof.sty + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_isar.eps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_isar.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/Thy/Functions.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/Functions.thy Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,1264 @@ +(* Title: doc-src/IsarAdvanced/Functions/Thy/Fundefs.thy + Author: Alexander Krauss, TU Muenchen + +Tutorial for function definitions with the new "function" package. +*) + +theory Functions +imports Main +begin + +section {* Function Definitions for Dummies *} + +text {* + In most cases, defining a recursive function is just as simple as other definitions: +*} + +fun fib :: "nat \ nat" +where + "fib 0 = 1" +| "fib (Suc 0) = 1" +| "fib (Suc (Suc n)) = fib n + fib (Suc n)" + +text {* + The syntax is rather self-explanatory: We introduce a function by + giving its name, its type, + and a set of defining recursive equations. + If we leave out the type, the most general type will be + inferred, which can sometimes lead to surprises: Since both @{term + "1::nat"} and @{text "+"} are overloaded, we would end up + with @{text "fib :: nat \ 'a::{one,plus}"}. +*} + +text {* + The function always terminates, since its argument gets smaller in + every recursive call. + Since HOL is a logic of total functions, termination is a + fundamental requirement to prevent inconsistencies\footnote{From the + \qt{definition} @{text "f(n) = f(n) + 1"} we could prove + @{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.}. + Isabelle tries to prove termination automatically when a definition + is made. In \S\ref{termination}, we will look at cases where this + fails and see what to do then. +*} + +subsection {* Pattern matching *} + +text {* \label{patmatch} + Like in functional programming, we can use pattern matching to + define functions. At the moment we will only consider \emph{constructor + patterns}, which only consist of datatype constructors and + variables. Furthermore, patterns must be linear, i.e.\ all variables + on the left hand side of an equation must be distinct. In + \S\ref{genpats} we discuss more general pattern matching. + + If patterns overlap, the order of the equations is taken into + account. The following function inserts a fixed element between any + two elements of a list: +*} + +fun sep :: "'a \ 'a list \ 'a list" +where + "sep a (x#y#xs) = x # a # sep a (y # xs)" +| "sep a xs = xs" + +text {* + Overlapping patterns are interpreted as \qt{increments} to what is + already there: The second equation is only meant for the cases where + the first one does not match. Consequently, Isabelle replaces it + internally by the remaining cases, making the patterns disjoint: +*} + +thm sep.simps + +text {* @{thm [display] sep.simps[no_vars]} *} + +text {* + \noindent The equations from function definitions are automatically used in + simplification: +*} + +lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]" +by simp + +subsection {* Induction *} + +text {* + + Isabelle provides customized induction rules for recursive + functions. These rules follow the recursive structure of the + definition. Here is the rule @{text sep.induct} arising from the + above definition of @{const sep}: + + @{thm [display] sep.induct} + + We have a step case for list with at least two elements, and two + base cases for the zero- and the one-element list. Here is a simple + proof about @{const sep} and @{const map} +*} + +lemma "map f (sep x ys) = sep (f x) (map f ys)" +apply (induct x ys rule: sep.induct) + +txt {* + We get three cases, like in the definition. + + @{subgoals [display]} +*} + +apply auto +done +text {* + + With the \cmd{fun} command, you can define about 80\% of the + functions that occur in practice. The rest of this tutorial explains + the remaining 20\%. +*} + + +section {* fun vs.\ function *} + +text {* + The \cmd{fun} command provides a + convenient shorthand notation for simple function definitions. In + this mode, Isabelle tries to solve all the necessary proof obligations + automatically. If any proof fails, the definition is + rejected. This can either mean that the definition is indeed faulty, + or that the default proof procedures are just not smart enough (or + rather: not designed) to handle the definition. + + By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or + solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: + +\end{isamarkuptext} + + +\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} +\cmd{fun} @{text "f :: \"}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\vspace*{6pt} +\end{minipage}\right] +\quad\equiv\quad +\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} +\cmd{function} @{text "("}\cmd{sequential}@{text ") f :: \"}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\\% +\cmd{by} @{text "pat_completeness auto"}\\% +\cmd{termination by} @{text "lexicographic_order"}\vspace{6pt} +\end{minipage} +\right]\] + +\begin{isamarkuptext} + \vspace*{1em} + \noindent Some details have now become explicit: + + \begin{enumerate} + \item The \cmd{sequential} option enables the preprocessing of + pattern overlaps which we already saw. Without this option, the equations + must already be disjoint and complete. The automatic completion only + works with constructor patterns. + + \item A function definition produces a proof obligation which + expresses completeness and compatibility of patterns (we talk about + this later). The combination of the methods @{text "pat_completeness"} and + @{text "auto"} is used to solve this proof obligation. + + \item A termination proof follows the definition, started by the + \cmd{termination} command. This will be explained in \S\ref{termination}. + \end{enumerate} + Whenever a \cmd{fun} command fails, it is usually a good idea to + expand the syntax to the more verbose \cmd{function} form, to see + what is actually going on. + *} + + +section {* Termination *} + +text {*\label{termination} + The method @{text "lexicographic_order"} is the default method for + termination proofs. It can prove termination of a + certain class of functions by searching for a suitable lexicographic + combination of size measures. Of course, not all functions have such + a simple termination argument. For them, we can specify the termination + relation manually. +*} + +subsection {* The {\tt relation} method *} +text{* + Consider the following function, which sums up natural numbers up to + @{text "N"}, using a counter @{text "i"}: +*} + +function sum :: "nat \ nat \ nat" +where + "sum i N = (if i > N then 0 else i + sum (Suc i) N)" +by pat_completeness auto + +text {* + \noindent The @{text "lexicographic_order"} method fails on this example, because none of the + arguments decreases in the recursive call, with respect to the standard size ordering. + To prove termination manually, we must provide a custom wellfounded relation. + + The termination argument for @{text "sum"} is based on the fact that + the \emph{difference} between @{text "i"} and @{text "N"} gets + smaller in every step, and that the recursion stops when @{text "i"} + is greater than @{text "N"}. Phrased differently, the expression + @{text "N + 1 - i"} always decreases. + + We can use this expression as a measure function suitable to prove termination. +*} + +termination sum +apply (relation "measure (\(i,N). N + 1 - i)") + +txt {* + The \cmd{termination} command sets up the termination goal for the + specified function @{text "sum"}. If the function name is omitted, it + implicitly refers to the last function definition. + + The @{text relation} method takes a relation of + type @{typ "('a \ 'a) set"}, where @{typ "'a"} is the argument type of + the function. If the function has multiple curried arguments, then + these are packed together into a tuple, as it happened in the above + example. + + The predefined function @{term[source] "measure :: ('a \ nat) \ ('a \ 'a) set"} constructs a + wellfounded relation from a mapping into the natural numbers (a + \emph{measure function}). + + After the invocation of @{text "relation"}, we must prove that (a) + the relation we supplied is wellfounded, and (b) that the arguments + of recursive calls indeed decrease with respect to the + relation: + + @{subgoals[display,indent=0]} + + These goals are all solved by @{text "auto"}: +*} + +apply auto +done + +text {* + Let us complicate the function a little, by adding some more + recursive calls: +*} + +function foo :: "nat \ nat \ nat" +where + "foo i N = (if i > N + then (if N = 0 then 0 else foo 0 (N - 1)) + else i + foo (Suc i) N)" +by pat_completeness auto + +text {* + When @{text "i"} has reached @{text "N"}, it starts at zero again + and @{text "N"} is decremented. + This corresponds to a nested + loop where one index counts up and the other down. Termination can + be proved using a lexicographic combination of two measures, namely + the value of @{text "N"} and the above difference. The @{const + "measures"} combinator generalizes @{text "measure"} by taking a + list of measure functions. +*} + +termination +by (relation "measures [\(i, N). N, \(i,N). N + 1 - i]") auto + +subsection {* How @{text "lexicographic_order"} works *} + +(*fun fails :: "nat \ nat list \ nat" +where + "fails a [] = a" +| "fails a (x#xs) = fails (x + a) (x # xs)" +*) + +text {* + To see how the automatic termination proofs work, let's look at an + example where it fails\footnote{For a detailed discussion of the + termination prover, see \cite{bulwahnKN07}}: + +\end{isamarkuptext} +\cmd{fun} @{text "fails :: \"nat \ nat list \ nat\""}\\% +\cmd{where}\\% +\hspace*{2ex}@{text "\"fails a [] = a\""}\\% +|\hspace*{1.5ex}@{text "\"fails a (x#xs) = fails (x + a) (x#xs)\""}\\ +\begin{isamarkuptext} + +\noindent Isabelle responds with the following error: + +\begin{isabelle} +*** Unfinished subgoals:\newline +*** (a, 1, <):\newline +*** \ 1.~@{text "\x. x = 0"}\newline +*** (a, 1, <=):\newline +*** \ 1.~False\newline +*** (a, 2, <):\newline +*** \ 1.~False\newline +*** Calls:\newline +*** a) @{text "(a, x # xs) -->> (x + a, x # xs)"}\newline +*** Measures:\newline +*** 1) @{text "\x. size (fst x)"}\newline +*** 2) @{text "\x. size (snd x)"}\newline +*** Result matrix:\newline +*** \ \ \ \ 1\ \ 2 \newline +*** a: ? <= \newline +*** Could not find lexicographic termination order.\newline +*** At command "fun".\newline +\end{isabelle} +*} + + +text {* + The key to this error message is the matrix at the bottom. The rows + of that matrix correspond to the different recursive calls (In our + case, there is just one). The columns are the function's arguments + (expressed through different measure functions, which map the + argument tuple to a natural number). + + The contents of the matrix summarize what is known about argument + descents: The second argument has a weak descent (@{text "<="}) at the + recursive call, and for the first argument nothing could be proved, + which is expressed by @{text "?"}. In general, there are the values + @{text "<"}, @{text "<="} and @{text "?"}. + + For the failed proof attempts, the unfinished subgoals are also + printed. Looking at these will often point to a missing lemma. + +% As a more real example, here is quicksort: +*} +(* +function qs :: "nat list \ nat list" +where + "qs [] = []" +| "qs (x#xs) = qs [y\xs. y < x] @ x # qs [y\xs. y \ x]" +by pat_completeness auto + +termination apply lexicographic_order + +text {* If we try @{text "lexicographic_order"} method, we get the + following error *} +termination by (lexicographic_order simp:l2) + +lemma l: "x \ y \ x < Suc y" by arith + +function + +*) + +section {* Mutual Recursion *} + +text {* + If two or more functions call one another mutually, they have to be defined + in one step. Here are @{text "even"} and @{text "odd"}: +*} + +function even :: "nat \ bool" + and odd :: "nat \ bool" +where + "even 0 = True" +| "odd 0 = False" +| "even (Suc n) = odd n" +| "odd (Suc n) = even n" +by pat_completeness auto + +text {* + To eliminate the mutual dependencies, Isabelle internally + creates a single function operating on the sum + type @{typ "nat + nat"}. Then, @{const even} and @{const odd} are + defined as projections. Consequently, termination has to be proved + simultaneously for both functions, by specifying a measure on the + sum type: +*} + +termination +by (relation "measure (\x. case x of Inl n \ n | Inr n \ n)") auto + +text {* + We could also have used @{text lexicographic_order}, which + supports mutual recursive termination proofs to a certain extent. +*} + +subsection {* Induction for mutual recursion *} + +text {* + + When functions are mutually recursive, proving properties about them + generally requires simultaneous induction. The induction rule @{text "even_odd.induct"} + generated from the above definition reflects this. + + Let us prove something about @{const even} and @{const odd}: +*} + +lemma even_odd_mod2: + "even n = (n mod 2 = 0)" + "odd n = (n mod 2 = 1)" + +txt {* + We apply simultaneous induction, specifying the induction variable + for both goals, separated by \cmd{and}: *} + +apply (induct n and n rule: even_odd.induct) + +txt {* + We get four subgoals, which correspond to the clauses in the + definition of @{const even} and @{const odd}: + @{subgoals[display,indent=0]} + Simplification solves the first two goals, leaving us with two + statements about the @{text "mod"} operation to prove: +*} + +apply simp_all + +txt {* + @{subgoals[display,indent=0]} + + \noindent These can be handled by Isabelle's arithmetic decision procedures. + +*} + +apply arith +apply arith +done + +text {* + In proofs like this, the simultaneous induction is really essential: + Even if we are just interested in one of the results, the other + one is necessary to strengthen the induction hypothesis. If we leave + out the statement about @{const odd} and just write @{term True} instead, + the same proof fails: +*} + +lemma failed_attempt: + "even n = (n mod 2 = 0)" + "True" +apply (induct n rule: even_odd.induct) + +txt {* + \noindent Now the third subgoal is a dead end, since we have no + useful induction hypothesis available: + + @{subgoals[display,indent=0]} +*} + +oops + +section {* General pattern matching *} +text{*\label{genpats} *} + +subsection {* Avoiding automatic pattern splitting *} + +text {* + + Up to now, we used pattern matching only on datatypes, and the + patterns were always disjoint and complete, and if they weren't, + they were made disjoint automatically like in the definition of + @{const "sep"} in \S\ref{patmatch}. + + This automatic splitting can significantly increase the number of + equations involved, and this is not always desirable. The following + example shows the problem: + + Suppose we are modeling incomplete knowledge about the world by a + three-valued datatype, which has values @{term "T"}, @{term "F"} + and @{term "X"} for true, false and uncertain propositions, respectively. +*} + +datatype P3 = T | F | X + +text {* \noindent Then the conjunction of such values can be defined as follows: *} + +fun And :: "P3 \ P3 \ P3" +where + "And T p = p" +| "And p T = p" +| "And p F = F" +| "And F p = F" +| "And X X = X" + + +text {* + This definition is useful, because the equations can directly be used + as simplification rules. But the patterns overlap: For example, + the expression @{term "And T T"} is matched by both the first and + the second equation. By default, Isabelle makes the patterns disjoint by + splitting them up, producing instances: +*} + +thm And.simps + +text {* + @{thm[indent=4] And.simps} + + \vspace*{1em} + \noindent There are several problems with this: + + \begin{enumerate} + \item If the datatype has many constructors, there can be an + explosion of equations. For @{const "And"}, we get seven instead of + five equations, which can be tolerated, but this is just a small + example. + + \item Since splitting makes the equations \qt{less general}, they + do not always match in rewriting. While the term @{term "And x F"} + can be simplified to @{term "F"} with the original equations, a + (manual) case split on @{term "x"} is now necessary. + + \item The splitting also concerns the induction rule @{text + "And.induct"}. Instead of five premises it now has seven, which + means that our induction proofs will have more cases. + + \item In general, it increases clarity if we get the same definition + back which we put in. + \end{enumerate} + + If we do not want the automatic splitting, we can switch it off by + leaving out the \cmd{sequential} option. However, we will have to + prove that our pattern matching is consistent\footnote{This prevents + us from defining something like @{term "f x = True"} and @{term "f x + = False"} simultaneously.}: +*} + +function And2 :: "P3 \ P3 \ P3" +where + "And2 T p = p" +| "And2 p T = p" +| "And2 p F = F" +| "And2 F p = F" +| "And2 X X = X" + +txt {* + \noindent Now let's look at the proof obligations generated by a + function definition. In this case, they are: + + @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} + + The first subgoal expresses the completeness of the patterns. It has + the form of an elimination rule and states that every @{term x} of + the function's input type must match at least one of the patterns\footnote{Completeness could + be equivalently stated as a disjunction of existential statements: +@{term "(\p. x = (T, p)) \ (\p. x = (p, T)) \ (\p. x = (p, F)) \ + (\p. x = (F, p)) \ (x = (X, X))"}, and you can use the method @{text atomize_elim} to get that form instead.}. If the patterns just involve + datatypes, we can solve it with the @{text "pat_completeness"} + method: +*} + +apply pat_completeness + +txt {* + The remaining subgoals express \emph{pattern compatibility}. We do + allow that an input value matches multiple patterns, but in this + case, the result (i.e.~the right hand sides of the equations) must + also be equal. For each pair of two patterns, there is one such + subgoal. Usually this needs injectivity of the constructors, which + is used automatically by @{text "auto"}. +*} + +by auto + + +subsection {* Non-constructor patterns *} + +text {* + Most of Isabelle's basic types take the form of inductive datatypes, + and usually pattern matching works on the constructors of such types. + However, this need not be always the case, and the \cmd{function} + command handles other kind of patterns, too. + + One well-known instance of non-constructor patterns are + so-called \emph{$n+k$-patterns}, which are a little controversial in + the functional programming world. Here is the initial fibonacci + example with $n+k$-patterns: +*} + +function fib2 :: "nat \ nat" +where + "fib2 0 = 1" +| "fib2 1 = 1" +| "fib2 (n + 2) = fib2 n + fib2 (Suc n)" + +(*<*)ML_val "goals_limit := 1"(*>*) +txt {* + This kind of matching is again justified by the proof of pattern + completeness and compatibility. + The proof obligation for pattern completeness states that every natural number is + either @{term "0::nat"}, @{term "1::nat"} or @{term "n + + (2::nat)"}: + + @{subgoals[display,indent=0]} + + This is an arithmetic triviality, but unfortunately the + @{text arith} method cannot handle this specific form of an + elimination rule. However, we can use the method @{text + "atomize_elim"} to do an ad-hoc conversion to a disjunction of + existentials, which can then be solved by the arithmetic decision procedure. + Pattern compatibility and termination are automatic as usual. +*} +(*<*)ML_val "goals_limit := 10"(*>*) +apply atomize_elim +apply arith +apply auto +done +termination by lexicographic_order +text {* + We can stretch the notion of pattern matching even more. The + following function is not a sensible functional program, but a + perfectly valid mathematical definition: +*} + +function ev :: "nat \ bool" +where + "ev (2 * n) = True" +| "ev (2 * n + 1) = False" +apply atomize_elim +by arith+ +termination by (relation "{}") simp + +text {* + This general notion of pattern matching gives you a certain freedom + in writing down specifications. However, as always, such freedom should + be used with care: + + If we leave the area of constructor + patterns, we have effectively departed from the world of functional + programming. This means that it is no longer possible to use the + code generator, and expect it to generate ML code for our + definitions. Also, such a specification might not work very well together with + simplification. Your mileage may vary. +*} + + +subsection {* Conditional equations *} + +text {* + The function package also supports conditional equations, which are + similar to guards in a language like Haskell. Here is Euclid's + algorithm written with conditional patterns\footnote{Note that the + patterns are also overlapping in the base case}: +*} + +function gcd :: "nat \ nat \ nat" +where + "gcd x 0 = x" +| "gcd 0 y = y" +| "x < y \ gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)" +| "\ x < y \ gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)" +by (atomize_elim, auto, arith) +termination by lexicographic_order + +text {* + By now, you can probably guess what the proof obligations for the + pattern completeness and compatibility look like. + + Again, functions with conditional patterns are not supported by the + code generator. +*} + + +subsection {* Pattern matching on strings *} + +text {* + As strings (as lists of characters) are normal datatypes, pattern + matching on them is possible, but somewhat problematic. Consider the + following definition: + +\end{isamarkuptext} +\noindent\cmd{fun} @{text "check :: \"string \ bool\""}\\% +\cmd{where}\\% +\hspace*{2ex}@{text "\"check (''good'') = True\""}\\% +@{text "| \"check s = False\""} +\begin{isamarkuptext} + + \noindent An invocation of the above \cmd{fun} command does not + terminate. What is the problem? Strings are lists of characters, and + characters are a datatype with a lot of constructors. Splitting the + catch-all pattern thus leads to an explosion of cases, which cannot + be handled by Isabelle. + + There are two things we can do here. Either we write an explicit + @{text "if"} on the right hand side, or we can use conditional patterns: +*} + +function check :: "string \ bool" +where + "check (''good'') = True" +| "s \ ''good'' \ check s = False" +by auto + + +section {* Partiality *} + +text {* + In HOL, all functions are total. A function @{term "f"} applied to + @{term "x"} always has the value @{term "f x"}, and there is no notion + of undefinedness. + This is why we have to do termination + proofs when defining functions: The proof justifies that the + function can be defined by wellfounded recursion. + + However, the \cmd{function} package does support partiality to a + certain extent. Let's look at the following function which looks + for a zero of a given function f. +*} + +function (*<*)(domintros, tailrec)(*>*)findzero :: "(nat \ nat) \ nat \ nat" +where + "findzero f n = (if f n = 0 then n else findzero f (Suc n))" +by pat_completeness auto +(*<*)declare findzero.simps[simp del](*>*) + +text {* + \noindent Clearly, any attempt of a termination proof must fail. And without + that, we do not get the usual rules @{text "findzero.simps"} and + @{text "findzero.induct"}. So what was the definition good for at all? +*} + +subsection {* Domain predicates *} + +text {* + The trick is that Isabelle has not only defined the function @{const findzero}, but also + a predicate @{term "findzero_dom"} that characterizes the values where the function + terminates: the \emph{domain} of the function. If we treat a + partial function just as a total function with an additional domain + predicate, we can derive simplification and + induction rules as we do for total functions. They are guarded + by domain conditions and are called @{text psimps} and @{text + pinduct}: +*} + +text {* + \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage} + \hfill(@{text "findzero.psimps"}) + \vspace{1em} + + \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage} + \hfill(@{text "findzero.pinduct"}) +*} + +text {* + Remember that all we + are doing here is use some tricks to make a total function appear + as if it was partial. We can still write the term @{term "findzero + (\x. 1) 0"} and like any other term of type @{typ nat} it is equal + to some natural number, although we might not be able to find out + which one. The function is \emph{underdefined}. + + But it is defined enough to prove something interesting about it. We + can prove that if @{term "findzero f n"} + terminates, it indeed returns a zero of @{term f}: +*} + +lemma findzero_zero: "findzero_dom (f, n) \ f (findzero f n) = 0" + +txt {* \noindent We apply induction as usual, but using the partial induction + rule: *} + +apply (induct f n rule: findzero.pinduct) + +txt {* \noindent This gives the following subgoals: + + @{subgoals[display,indent=0]} + + \noindent The hypothesis in our lemma was used to satisfy the first premise in + the induction rule. However, we also get @{term + "findzero_dom (f, n)"} as a local assumption in the induction step. This + allows to unfold @{term "findzero f n"} using the @{text psimps} + rule, and the rest is trivial. Since the @{text psimps} rules carry the + @{text "[simp]"} attribute by default, we just need a single step: + *} +apply simp +done + +text {* + Proofs about partial functions are often not harder than for total + functions. Fig.~\ref{findzero_isar} shows a slightly more + complicated proof written in Isar. It is verbose enough to show how + partiality comes into play: From the partial induction, we get an + additional domain condition hypothesis. Observe how this condition + is applied when calls to @{term findzero} are unfolded. +*} + +text_raw {* +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +*} +lemma "\findzero_dom (f, n); x \ {n ..< findzero f n}\ \ f x \ 0" +proof (induct rule: findzero.pinduct) + fix f n assume dom: "findzero_dom (f, n)" + and IH: "\f n \ 0; x \ {Suc n ..< findzero f (Suc n)}\ \ f x \ 0" + and x_range: "x \ {n ..< findzero f n}" + have "f n \ 0" + proof + assume "f n = 0" + with dom have "findzero f n = n" by simp + with x_range show False by auto + qed + + from x_range have "x = n \ x \ {Suc n ..< findzero f n}" by auto + thus "f x \ 0" + proof + assume "x = n" + with `f n \ 0` show ?thesis by simp + next + assume "x \ {Suc n ..< findzero f n}" + with dom and `f n \ 0` have "x \ {Suc n ..< findzero f (Suc n)}" by simp + with IH and `f n \ 0` + show ?thesis by simp + qed +qed +text_raw {* +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{A proof about a partial function}\label{findzero_isar} +\end{figure} +*} + +subsection {* Partial termination proofs *} + +text {* + Now that we have proved some interesting properties about our + function, we should turn to the domain predicate and see if it is + actually true for some values. Otherwise we would have just proved + lemmas with @{term False} as a premise. + + Essentially, we need some introduction rules for @{text + findzero_dom}. The function package can prove such domain + introduction rules automatically. But since they are not used very + often (they are almost never needed if the function is total), this + functionality is disabled by default for efficiency reasons. So we have to go + back and ask for them explicitly by passing the @{text + "(domintros)"} option to the function package: + +\vspace{1ex} +\noindent\cmd{function} @{text "(domintros) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% +\cmd{where}\isanewline% +\ \ \ldots\\ + + \noindent Now the package has proved an introduction rule for @{text findzero_dom}: +*} + +thm findzero.domintros + +text {* + @{thm[display] findzero.domintros} + + Domain introduction rules allow to show that a given value lies in the + domain of a function, if the arguments of all recursive calls + are in the domain as well. They allow to do a \qt{single step} in a + termination proof. Usually, you want to combine them with a suitable + induction principle. + + Since our function increases its argument at recursive calls, we + need an induction principle which works \qt{backwards}. We will use + @{text inc_induct}, which allows to do induction from a fixed number + \qt{downwards}: + + \begin{center}@{thm inc_induct}\hfill(@{text "inc_induct"})\end{center} + + Figure \ref{findzero_term} gives a detailed Isar proof of the fact + that @{text findzero} terminates if there is a zero which is greater + or equal to @{term n}. First we derive two useful rules which will + solve the base case and the step case of the induction. The + induction is then straightforward, except for the unusual induction + principle. + +*} + +text_raw {* +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +*} +lemma findzero_termination: + assumes "x \ n" and "f x = 0" + shows "findzero_dom (f, n)" +proof - + have base: "findzero_dom (f, x)" + by (rule findzero.domintros) (simp add:`f x = 0`) + + have step: "\i. findzero_dom (f, Suc i) + \ findzero_dom (f, i)" + by (rule findzero.domintros) simp + + from `x \ n` show ?thesis + proof (induct rule:inc_induct) + show "findzero_dom (f, x)" by (rule base) + next + fix i assume "findzero_dom (f, Suc i)" + thus "findzero_dom (f, i)" by (rule step) + qed +qed +text_raw {* +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{Termination proof for @{text findzero}}\label{findzero_term} +\end{figure} +*} + +text {* + Again, the proof given in Fig.~\ref{findzero_term} has a lot of + detail in order to explain the principles. Using more automation, we + can also have a short proof: +*} + +lemma findzero_termination_short: + assumes zero: "x >= n" + assumes [simp]: "f x = 0" + shows "findzero_dom (f, n)" +using zero +by (induct rule:inc_induct) (auto intro: findzero.domintros) + +text {* + \noindent It is simple to combine the partial correctness result with the + termination lemma: +*} + +lemma findzero_total_correctness: + "f x = 0 \ f (findzero f 0) = 0" +by (blast intro: findzero_zero findzero_termination) + +subsection {* Definition of the domain predicate *} + +text {* + Sometimes it is useful to know what the definition of the domain + predicate looks like. Actually, @{text findzero_dom} is just an + abbreviation: + + @{abbrev[display] findzero_dom} + + The domain predicate is the \emph{accessible part} of a relation @{const + findzero_rel}, which was also created internally by the function + package. @{const findzero_rel} is just a normal + inductive predicate, so we can inspect its definition by + looking at the introduction rules @{text findzero_rel.intros}. + In our case there is just a single rule: + + @{thm[display] findzero_rel.intros} + + The predicate @{const findzero_rel} + describes the \emph{recursion relation} of the function + definition. The recursion relation is a binary relation on + the arguments of the function that relates each argument to its + recursive calls. In general, there is one introduction rule for each + recursive call. + + The predicate @{term "accp findzero_rel"} is the accessible part of + that relation. An argument belongs to the accessible part, if it can + be reached in a finite number of steps (cf.~its definition in @{text + "Wellfounded.thy"}). + + Since the domain predicate is just an abbreviation, you can use + lemmas for @{const accp} and @{const findzero_rel} directly. Some + lemmas which are occasionally useful are @{text accpI}, @{text + accp_downward}, and of course the introduction and elimination rules + for the recursion relation @{text "findzero.intros"} and @{text "findzero.cases"}. +*} + +(*lemma findzero_nicer_domintros: + "f x = 0 \ findzero_dom (f, x)" + "findzero_dom (f, Suc x) \ findzero_dom (f, x)" +by (rule accpI, erule findzero_rel.cases, auto)+ +*) + +subsection {* A Useful Special Case: Tail recursion *} + +text {* + The domain predicate is our trick that allows us to model partiality + in a world of total functions. The downside of this is that we have + to carry it around all the time. The termination proof above allowed + us to replace the abstract @{term "findzero_dom (f, n)"} by the more + concrete @{term "(x \ n \ f x = (0::nat))"}, but the condition is still + there and can only be discharged for special cases. + In particular, the domain predicate guards the unfolding of our + function, since it is there as a condition in the @{text psimp} + rules. + + Now there is an important special case: We can actually get rid + of the condition in the simplification rules, \emph{if the function + is tail-recursive}. The reason is that for all tail-recursive + equations there is a total function satisfying them, even if they + are non-terminating. + +% A function is tail recursive, if each call to the function is either +% equal +% +% So the outer form of the +% +%if it can be written in the following +% form: +% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} + + + The function package internally does the right construction and can + derive the unconditional simp rules, if we ask it to do so. Luckily, + our @{const "findzero"} function is tail-recursive, so we can just go + back and add another option to the \cmd{function} command: + +\vspace{1ex} +\noindent\cmd{function} @{text "(domintros, tailrec) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% +\cmd{where}\isanewline% +\ \ \ldots\\% + + + \noindent Now, we actually get unconditional simplification rules, even + though the function is partial: +*} + +thm findzero.simps + +text {* + @{thm[display] findzero.simps} + + \noindent Of course these would make the simplifier loop, so we better remove + them from the simpset: +*} + +declare findzero.simps[simp del] + +text {* + Getting rid of the domain conditions in the simplification rules is + not only useful because it simplifies proofs. It is also required in + order to use Isabelle's code generator to generate ML code + from a function definition. + Since the code generator only works with equations, it cannot be + used with @{text "psimp"} rules. Thus, in order to generate code for + partial functions, they must be defined as a tail recursion. + Luckily, many functions have a relatively natural tail recursive + definition. +*} + +section {* Nested recursion *} + +text {* + Recursive calls which are nested in one another frequently cause + complications, since their termination proof can depend on a partial + correctness property of the function itself. + + As a small example, we define the \qt{nested zero} function: +*} + +function nz :: "nat \ nat" +where + "nz 0 = 0" +| "nz (Suc n) = nz (nz n)" +by pat_completeness auto + +text {* + If we attempt to prove termination using the identity measure on + naturals, this fails: +*} + +termination + apply (relation "measure (\n. n)") + apply auto + +txt {* + We get stuck with the subgoal + + @{subgoals[display]} + + Of course this statement is true, since we know that @{const nz} is + the zero function. And in fact we have no problem proving this + property by induction. +*} +(*<*)oops(*>*) +lemma nz_is_zero: "nz_dom n \ nz n = 0" + by (induct rule:nz.pinduct) auto + +text {* + We formulate this as a partial correctness lemma with the condition + @{term "nz_dom n"}. This allows us to prove it with the @{text + pinduct} rule before we have proved termination. With this lemma, + the termination proof works as expected: +*} + +termination + by (relation "measure (\n. n)") (auto simp: nz_is_zero) + +text {* + As a general strategy, one should prove the statements needed for + termination as a partial property first. Then they can be used to do + the termination proof. This also works for less trivial + examples. Figure \ref{f91} defines the 91-function, a well-known + challenge problem due to John McCarthy, and proves its termination. +*} + +text_raw {* +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +*} + +function f91 :: "nat \ nat" +where + "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" +by pat_completeness auto + +lemma f91_estimate: + assumes trm: "f91_dom n" + shows "n < f91 n + 11" +using trm by induct auto + +termination +proof + let ?R = "measure (\x. 101 - x)" + show "wf ?R" .. + + fix n :: nat assume "\ 100 < n" -- "Assumptions for both calls" + + thus "(n + 11, n) \ ?R" by simp -- "Inner call" + + assume inner_trm: "f91_dom (n + 11)" -- "Outer call" + with f91_estimate have "n + 11 < f91 (n + 11) + 11" . + with `\ 100 < n` show "(f91 (n + 11), n) \ ?R" by simp +qed + +text_raw {* +\isamarkupfalse\isabellestyle{tt} +\end{minipage} +\vspace{6pt}\hrule +\caption{McCarthy's 91-function}\label{f91} +\end{figure} +*} + + +section {* Higher-Order Recursion *} + +text {* + Higher-order recursion occurs when recursive calls + are passed as arguments to higher-order combinators such as @{const + map}, @{term filter} etc. + As an example, imagine a datatype of n-ary trees: +*} + +datatype 'a tree = + Leaf 'a +| Branch "'a tree list" + + +text {* \noindent We can define a function which swaps the left and right subtrees recursively, using the + list functions @{const rev} and @{const map}: *} + +fun mirror :: "'a tree \ 'a tree" +where + "mirror (Leaf n) = Leaf n" +| "mirror (Branch l) = Branch (rev (map mirror l))" + +text {* + Although the definition is accepted without problems, let us look at the termination proof: +*} + +termination proof + txt {* + + As usual, we have to give a wellfounded relation, such that the + arguments of the recursive calls get smaller. But what exactly are + the arguments of the recursive calls when mirror is given as an + argument to @{const map}? Isabelle gives us the + subgoals + + @{subgoals[display,indent=0]} + + So the system seems to know that @{const map} only + applies the recursive call @{term "mirror"} to elements + of @{term "l"}, which is essential for the termination proof. + + This knowledge about @{const map} is encoded in so-called congruence rules, + which are special theorems known to the \cmd{function} command. The + rule for @{const map} is + + @{thm[display] map_cong} + + You can read this in the following way: Two applications of @{const + map} are equal, if the list arguments are equal and the functions + coincide on the elements of the list. This means that for the value + @{term "map f l"} we only have to know how @{term f} behaves on + the elements of @{term l}. + + Usually, one such congruence rule is + needed for each higher-order construct that is used when defining + new functions. In fact, even basic functions like @{const + If} and @{const Let} are handled by this mechanism. The congruence + rule for @{const If} states that the @{text then} branch is only + relevant if the condition is true, and the @{text else} branch only if it + is false: + + @{thm[display] if_cong} + + Congruence rules can be added to the + function package by giving them the @{term fundef_cong} attribute. + + The constructs that are predefined in Isabelle, usually + come with the respective congruence rules. + But if you define your own higher-order functions, you may have to + state and prove the required congruence rules yourself, if you want to use your + functions in recursive definitions. +*} +(*<*)oops(*>*) + +subsection {* Congruence Rules and Evaluation Order *} + +text {* + Higher order logic differs from functional programming languages in + that it has no built-in notion of evaluation order. A program is + just a set of equations, and it is not specified how they must be + evaluated. + + However for the purpose of function definition, we must talk about + evaluation order implicitly, when we reason about termination. + Congruence rules express that a certain evaluation order is + consistent with the logical definition. + + Consider the following function. +*} + +function f :: "nat \ bool" +where + "f n = (n = 0 \ f (n - 1))" +(*<*)by pat_completeness auto(*>*) + +text {* + For this definition, the termination proof fails. The default configuration + specifies no congruence rule for disjunction. We have to add a + congruence rule that specifies left-to-right evaluation order: + + \vspace{1ex} + \noindent @{thm disj_cong}\hfill(@{text "disj_cong"}) + \vspace{1ex} + + Now the definition works without problems. Note how the termination + proof depends on the extra condition that we get from the congruence + rule. + + However, as evaluation is not a hard-wired concept, we + could just turn everything around by declaring a different + congruence rule. Then we can make the reverse definition: +*} + +lemma disj_cong2[fundef_cong]: + "(\ Q' \ P = P') \ (Q = Q') \ (P \ Q) = (P' \ Q')" + by blast + +fun f' :: "nat \ bool" +where + "f' n = (f' (n - 1) \ n = 0)" + +text {* + \noindent These examples show that, in general, there is no \qt{best} set of + congruence rules. + + However, such tweaking should rarely be necessary in + practice, as most of the time, the default set of congruence rules + works well. +*} + +end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/ROOT.ML Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,4 @@ + +(* $Id$ *) + +use_thy "Functions"; diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/Thy/document/Functions.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/document/Functions.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,1985 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Functions}% +% +\isadelimtheory +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Functions\isanewline +\isakeyword{imports}\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Function Definitions for Dummies% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In most cases, defining a recursive function is just as simple as other definitions:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ n\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +The syntax is rather self-explanatory: We introduce a function by + giving its name, its type, + and a set of defining recursive equations. + If we leave out the type, the most general type will be + inferred, which can sometimes lead to surprises: Since both \isa{{\isadigit{1}}} and \isa{{\isacharplus}} are overloaded, we would end up + with \isa{fib\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isacharcolon}{\isacharcolon}{\isacharbraceleft}one{\isacharcomma}plus{\isacharbraceright}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +The function always terminates, since its argument gets smaller in + every recursive call. + Since HOL is a logic of total functions, termination is a + fundamental requirement to prevent inconsistencies\footnote{From the + \qt{definition} \isa{f{\isacharparenleft}n{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}n{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}} we could prove + \isa{{\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}} by subtracting \isa{f{\isacharparenleft}n{\isacharparenright}} on both sides.}. + Isabelle tries to prove termination automatically when a definition + is made. In \S\ref{termination}, we will look at cases where this + fails and see what to do then.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Pattern matching% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{patmatch} + Like in functional programming, we can use pattern matching to + define functions. At the moment we will only consider \emph{constructor + patterns}, which only consist of datatype constructors and + variables. Furthermore, patterns must be linear, i.e.\ all variables + on the left hand side of an equation must be distinct. In + \S\ref{genpats} we discuss more general pattern matching. + + If patterns overlap, the order of the equations is taken into + account. The following function inserts a fixed element between any + two elements of a list:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}sep\ a\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}sep\ a\ xs\ \ \ \ \ \ \ {\isacharequal}\ xs{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Overlapping patterns are interpreted as \qt{increments} to what is + already there: The second equation is only meant for the cases where + the first one does not match. Consequently, Isabelle replaces it + internally by the remaining cases, making the patterns disjoint:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ sep{\isachardot}simps% +\begin{isamarkuptext}% +\begin{isabelle}% +sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\isasep\isanewline% +sep\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\isasep\isanewline% +sep\ a\ {\isacharbrackleft}v{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}v{\isacharbrackright}% +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent The equations from function definitions are automatically used in + simplification:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}sep\ {\isadigit{0}}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Induction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle provides customized induction rules for recursive + functions. These rules follow the recursive structure of the + definition. Here is the rule \isa{sep{\isachardot}induct} arising from the + above definition of \isa{sep}: + + \begin{isabelle}% +{\isasymlbrakk}{\isasymAnd}a\ x\ y\ xs{\isachardot}\ {\isacharquery}P\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}P\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharsemicolon}\ {\isasymAnd}a{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ {\isasymAnd}a\ v{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% +\end{isabelle} + + We have a step case for list with at least two elements, and two + base cases for the zero- and the one-element list. Here is a simple + proof about \isa{sep} and \isa{map}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}map\ f\ {\isacharparenleft}sep\ x\ ys{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}map\ f\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ x\ ys\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}% +\begin{isamarkuptxt}% +We get three cases, like in the definition. + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ x\ y\ xs{\isachardot}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ v{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}% +\end{isabelle}% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ auto\ \isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +With the \cmd{fun} command, you can define about 80\% of the + functions that occur in practice. The rest of this tutorial explains + the remaining 20\%.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{fun vs.\ function% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \cmd{fun} command provides a + convenient shorthand notation for simple function definitions. In + this mode, Isabelle tries to solve all the necessary proof obligations + automatically. If any proof fails, the definition is + rejected. This can either mean that the definition is indeed faulty, + or that the default proof procedures are just not smart enough (or + rather: not designed) to handle the definition. + + By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or + solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: + +\end{isamarkuptext} + + +\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} +\cmd{fun} \isa{f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\vspace*{6pt} +\end{minipage}\right] +\quad\equiv\quad +\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} +\cmd{function} \isa{{\isacharparenleft}}\cmd{sequential}\isa{{\isacharparenright}\ f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\\% +\cmd{by} \isa{pat{\isacharunderscore}completeness\ auto}\\% +\cmd{termination by} \isa{lexicographic{\isacharunderscore}order}\vspace{6pt} +\end{minipage} +\right]\] + +\begin{isamarkuptext} + \vspace*{1em} + \noindent Some details have now become explicit: + + \begin{enumerate} + \item The \cmd{sequential} option enables the preprocessing of + pattern overlaps which we already saw. Without this option, the equations + must already be disjoint and complete. The automatic completion only + works with constructor patterns. + + \item A function definition produces a proof obligation which + expresses completeness and compatibility of patterns (we talk about + this later). The combination of the methods \isa{pat{\isacharunderscore}completeness} and + \isa{auto} is used to solve this proof obligation. + + \item A termination proof follows the definition, started by the + \cmd{termination} command. This will be explained in \S\ref{termination}. + \end{enumerate} + Whenever a \cmd{fun} command fails, it is usually a good idea to + expand the syntax to the more verbose \cmd{function} form, to see + what is actually going on.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Termination% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{termination} + The method \isa{lexicographic{\isacharunderscore}order} is the default method for + termination proofs. It can prove termination of a + certain class of functions by searching for a suitable lexicographic + combination of size measures. Of course, not all functions have such + a simple termination argument. For them, we can specify the termination + relation manually.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{The {\tt relation} method% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Consider the following function, which sums up natural numbers up to + \isa{N}, using a counter \isa{i}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ sum\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}sum\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ then\ {\isadigit{0}}\ else\ i\ {\isacharplus}\ sum\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent The \isa{lexicographic{\isacharunderscore}order} method fails on this example, because none of the + arguments decreases in the recursive call, with respect to the standard size ordering. + To prove termination manually, we must provide a custom wellfounded relation. + + The termination argument for \isa{sum} is based on the fact that + the \emph{difference} between \isa{i} and \isa{N} gets + smaller in every step, and that the recursion stops when \isa{i} + is greater than \isa{N}. Phrased differently, the expression + \isa{N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i} always decreases. + + We can use this expression as a measure function suitable to prove termination.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\ sum\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% +\begin{isamarkuptxt}% +The \cmd{termination} command sets up the termination goal for the + specified function \isa{sum}. If the function name is omitted, it + implicitly refers to the last function definition. + + The \isa{relation} method takes a relation of + type \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}, where \isa{{\isacharprime}a} is the argument type of + the function. If the function has multiple curried arguments, then + these are packed together into a tuple, as it happened in the above + example. + + The predefined function \isa{{\isachardoublequote}measure\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}} constructs a + wellfounded relation from a mapping into the natural numbers (a + \emph{measure function}). + + After the invocation of \isa{relation}, we must prove that (a) + the relation we supplied is wellfounded, and (b) that the arguments + of recursive calls indeed decrease with respect to the + relation: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharparenleft}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}i\ N{\isachardot}\ {\isasymnot}\ N\ {\isacharless}\ i\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharparenleft}Suc\ i{\isacharcomma}\ N{\isacharparenright}{\isacharcomma}\ i{\isacharcomma}\ N{\isacharparenright}\ {\isasymin}\ measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}% +\end{isabelle} + + These goals are all solved by \isa{auto}:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ auto\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Let us complicate the function a little, by adding some more + recursive calls:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ foo\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}foo\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ \isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ then\ {\isacharparenleft}if\ N\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isadigit{0}}\ else\ foo\ {\isadigit{0}}\ {\isacharparenleft}N\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ i\ {\isacharplus}\ foo\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +When \isa{i} has reached \isa{N}, it starts at zero again + and \isa{N} is decremented. + This corresponds to a nested + loop where one index counts up and the other down. Termination can + be proved using a lexicographic combination of two measures, namely + the value of \isa{N} and the above difference. The \isa{measures} combinator generalizes \isa{measure} by taking a + list of measure functions.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\ \isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measures\ {\isacharbrackleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N{\isacharcomma}\ {\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{How \isa{lexicographic{\isacharunderscore}order} works% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +To see how the automatic termination proofs work, let's look at an + example where it fails\footnote{For a detailed discussion of the + termination prover, see \cite{bulwahnKN07}}: + +\end{isamarkuptext} +\cmd{fun} \isa{fails\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat\ list\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% +\cmd{where}\\% +\hspace*{2ex}\isa{{\isachardoublequote}fails\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ a{\isachardoublequote}}\\% +|\hspace*{1.5ex}\isa{{\isachardoublequote}fails\ a\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ fails\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharparenright}\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}{\isachardoublequote}}\\ +\begin{isamarkuptext} + +\noindent Isabelle responds with the following error: + +\begin{isabelle} +*** Unfinished subgoals:\newline +*** (a, 1, <):\newline +*** \ 1.~\isa{{\isasymAnd}x{\isachardot}\ x\ {\isacharequal}\ {\isadigit{0}}}\newline +*** (a, 1, <=):\newline +*** \ 1.~False\newline +*** (a, 2, <):\newline +*** \ 1.~False\newline +*** Calls:\newline +*** a) \isa{{\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharminus}{\isacharminus}{\isachargreater}{\isachargreater}\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}}\newline +*** Measures:\newline +*** 1) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}fst\ x{\isacharparenright}}\newline +*** 2) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}snd\ x{\isacharparenright}}\newline +*** Result matrix:\newline +*** \ \ \ \ 1\ \ 2 \newline +*** a: ? <= \newline +*** Could not find lexicographic termination order.\newline +*** At command "fun".\newline +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +The key to this error message is the matrix at the bottom. The rows + of that matrix correspond to the different recursive calls (In our + case, there is just one). The columns are the function's arguments + (expressed through different measure functions, which map the + argument tuple to a natural number). + + The contents of the matrix summarize what is known about argument + descents: The second argument has a weak descent (\isa{{\isacharless}{\isacharequal}}) at the + recursive call, and for the first argument nothing could be proved, + which is expressed by \isa{{\isacharquery}}. In general, there are the values + \isa{{\isacharless}}, \isa{{\isacharless}{\isacharequal}} and \isa{{\isacharquery}}. + + For the failed proof attempts, the unfinished subgoals are also + printed. Looking at these will often point to a missing lemma. + +% As a more real example, here is quicksort:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Mutual Recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +If two or more functions call one another mutually, they have to be defined + in one step. Here are \isa{even} and \isa{odd}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ even\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ odd\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}even\ {\isadigit{0}}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isadigit{0}}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ odd\ n{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ even\ n{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +To eliminate the mutual dependencies, Isabelle internally + creates a single function operating on the sum + type \isa{nat\ {\isacharplus}\ nat}. Then, \isa{even} and \isa{odd} are + defined as projections. Consequently, termination has to be proved + simultaneously for both functions, by specifying a measure on the + sum type:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\ \isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ case\ x\ of\ Inl\ n\ {\isasymRightarrow}\ n\ {\isacharbar}\ Inr\ n\ {\isasymRightarrow}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +We could also have used \isa{lexicographic{\isacharunderscore}order}, which + supports mutual recursive termination proofs to a certain extent.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Induction for mutual recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +When functions are mutually recursive, proving properties about them + generally requires simultaneous induction. The induction rule \isa{even{\isacharunderscore}odd{\isachardot}induct} + generated from the above definition reflects this. + + Let us prove something about \isa{even} and \isa{odd}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ even{\isacharunderscore}odd{\isacharunderscore}mod{\isadigit{2}}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +We apply simultaneous induction, specifying the induction variable + for both goals, separated by \cmd{and}:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ n\ \isakeyword{and}\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% +\begin{isamarkuptxt}% +We get four subgoals, which correspond to the clauses in the + definition of \isa{even} and \isa{odd}: + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ odd\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}% +\end{isabelle} + Simplification solves the first two goals, leaving us with two + statements about the \isa{mod} operation to prove:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\begin{isamarkuptxt}% +\begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}% +\end{isabelle} + + \noindent These can be handled by Isabelle's arithmetic decision procedures.% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ arith\isanewline +\isacommand{apply}\isamarkupfalse% +\ arith\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +In proofs like this, the simultaneous induction is really essential: + Even if we are just interested in one of the results, the other + one is necessary to strengthen the induction hypothesis. If we leave + out the statement about \isa{odd} and just write \isa{True} instead, + the same proof fails:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ failed{\isacharunderscore}attempt{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}True{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent Now the third subgoal is a dead end, since we have no + useful induction hypothesis available: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ True\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ True\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ True% +\end{isabelle}% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{oops}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsection{General pattern matching% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{genpats}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Avoiding automatic pattern splitting% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Up to now, we used pattern matching only on datatypes, and the + patterns were always disjoint and complete, and if they weren't, + they were made disjoint automatically like in the definition of + \isa{sep} in \S\ref{patmatch}. + + This automatic splitting can significantly increase the number of + equations involved, and this is not always desirable. The following + example shows the problem: + + Suppose we are modeling incomplete knowledge about the world by a + three-valued datatype, which has values \isa{T}, \isa{F} + and \isa{X} for true, false and uncertain propositions, respectively.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ P{\isadigit{3}}\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ X% +\begin{isamarkuptext}% +\noindent Then the conjunction of such values can be defined as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ And\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}And\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% +\begin{isamarkuptext}% +This definition is useful, because the equations can directly be used + as simplification rules. But the patterns overlap: For example, + the expression \isa{And\ T\ T} is matched by both the first and + the second equation. By default, Isabelle makes the patterns disjoint by + splitting them up, producing instances:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ And{\isachardot}simps% +\begin{isamarkuptext}% +\isa{And\ T\ {\isacharquery}p\ {\isacharequal}\ {\isacharquery}p\isasep\isanewline% +And\ F\ T\ {\isacharequal}\ F\isasep\isanewline% +And\ X\ T\ {\isacharequal}\ X\isasep\isanewline% +And\ F\ F\ {\isacharequal}\ F\isasep\isanewline% +And\ X\ F\ {\isacharequal}\ F\isasep\isanewline% +And\ F\ X\ {\isacharequal}\ F\isasep\isanewline% +And\ X\ X\ {\isacharequal}\ X} + + \vspace*{1em} + \noindent There are several problems with this: + + \begin{enumerate} + \item If the datatype has many constructors, there can be an + explosion of equations. For \isa{And}, we get seven instead of + five equations, which can be tolerated, but this is just a small + example. + + \item Since splitting makes the equations \qt{less general}, they + do not always match in rewriting. While the term \isa{And\ x\ F} + can be simplified to \isa{F} with the original equations, a + (manual) case split on \isa{x} is now necessary. + + \item The splitting also concerns the induction rule \isa{And{\isachardot}induct}. Instead of five premises it now has seven, which + means that our induction proofs will have more cases. + + \item In general, it increases clarity if we get the same definition + back which we put in. + \end{enumerate} + + If we do not want the automatic splitting, we can switch it off by + leaving out the \cmd{sequential} option. However, we will have to + prove that our pattern matching is consistent\footnote{This prevents + us from defining something like \isa{f\ x\ {\isacharequal}\ True} and \isa{f\ x\ {\isacharequal}\ False} simultaneously.}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ And{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}And{\isadigit{2}}\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +\noindent Now let's look at the proof obligations generated by a + function definition. In this case, they are: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ \ }{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ }{\isasymLongrightarrow}\ P\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{5}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X\isanewline +\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline +\ {\isadigit{8}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{9}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{1}}{\isadigit{0}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X% +\end{isabelle}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} + + The first subgoal expresses the completeness of the patterns. It has + the form of an elimination rule and states that every \isa{x} of + the function's input type must match at least one of the patterns\footnote{Completeness could + be equivalently stated as a disjunction of existential statements: +\isa{{\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}}, and you can use the method \isa{atomize{\isacharunderscore}elim} to get that form instead.}. If the patterns just involve + datatypes, we can solve it with the \isa{pat{\isacharunderscore}completeness} + method:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ pat{\isacharunderscore}completeness% +\begin{isamarkuptxt}% +The remaining subgoals express \emph{pattern compatibility}. We do + allow that an input value matches multiple patterns, but in this + case, the result (i.e.~the right hand sides of the equations) must + also be equal. For each pair of two patterns, there is one such + subgoal. Usually this needs injectivity of the constructors, which + is used automatically by \isa{auto}.% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Non-constructor patterns% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Most of Isabelle's basic types take the form of inductive datatypes, + and usually pattern matching works on the constructors of such types. + However, this need not be always the case, and the \cmd{function} + command handles other kind of patterns, too. + + One well-known instance of non-constructor patterns are + so-called \emph{$n+k$-patterns}, which are a little controversial in + the functional programming world. Here is the initial fibonacci + example with $n+k$-patterns:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ fib{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}\ n\ {\isacharplus}\ fib{\isadigit{2}}\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +This kind of matching is again justified by the proof of pattern + completeness and compatibility. + The proof obligation for pattern completeness states that every natural number is + either \isa{{\isadigit{0}}}, \isa{{\isadigit{1}}} or \isa{n\ {\isacharplus}\ {\isadigit{2}}}: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline +\ {\isadigit{5}}{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline +\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}n\ na{\isachardot}\isanewline +\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }n\ {\isacharplus}\ {\isadigit{2}}\ {\isacharequal}\ na\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\isanewline +\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ na\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ na{\isacharparenright}% +\end{isabelle} + + This is an arithmetic triviality, but unfortunately the + \isa{arith} method cannot handle this specific form of an + elimination rule. However, we can use the method \isa{atomize{\isacharunderscore}elim} to do an ad-hoc conversion to a disjunction of + existentials, which can then be solved by the arithmetic decision procedure. + Pattern compatibility and termination are automatic as usual.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ atomize{\isacharunderscore}elim\isanewline +\isacommand{apply}\isamarkupfalse% +\ arith\isanewline +\isacommand{apply}\isamarkupfalse% +\ auto\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ lexicographic{\isacharunderscore}order% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +We can stretch the notion of pattern matching even more. The + following function is not a sensible functional program, but a + perfectly valid mathematical definition:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ ev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ atomize{\isacharunderscore}elim\isanewline +\isacommand{by}\isamarkupfalse% +\ arith{\isacharplus}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}{\isacharparenright}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +This general notion of pattern matching gives you a certain freedom + in writing down specifications. However, as always, such freedom should + be used with care: + + If we leave the area of constructor + patterns, we have effectively departed from the world of functional + programming. This means that it is no longer possible to use the + code generator, and expect it to generate ML code for our + definitions. Also, such a specification might not work very well together with + simplification. Your mileage may vary.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Conditional equations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The function package also supports conditional equations, which are + similar to guards in a language like Haskell. Here is Euclid's + algorithm written with conditional patterns\footnote{Note that the + patterns are also overlapping in the base case}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}gcd\ x\ {\isadigit{0}}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}gcd\ {\isadigit{0}}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}y\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}{\isasymnot}\ x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}x\ {\isacharminus}\ y{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}atomize{\isacharunderscore}elim{\isacharcomma}\ auto{\isacharcomma}\ arith{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ lexicographic{\isacharunderscore}order% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +By now, you can probably guess what the proof obligations for the + pattern completeness and compatibility look like. + + Again, functions with conditional patterns are not supported by the + code generator.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Pattern matching on strings% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As strings (as lists of characters) are normal datatypes, pattern + matching on them is possible, but somewhat problematic. Consider the + following definition: + +\end{isamarkuptext} +\noindent\cmd{fun} \isa{check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}string\ {\isasymRightarrow}\ bool{\isachardoublequote}}\\% +\cmd{where}\\% +\hspace*{2ex}\isa{{\isachardoublequote}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}}\\% +\isa{{\isacharbar}\ {\isachardoublequote}check\ s\ {\isacharequal}\ False{\isachardoublequote}} +\begin{isamarkuptext} + + \noindent An invocation of the above \cmd{fun} command does not + terminate. What is the problem? Strings are lists of characters, and + characters are a datatype with a lot of constructors. Splitting the + catch-all pattern thus leads to an explosion of cases, which cannot + be handled by Isabelle. + + There are two things we can do here. Either we write an explicit + \isa{if} on the right hand side, or we can use conditional patterns:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}s\ {\isasymnoteq}\ {\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}\ {\isasymLongrightarrow}\ check\ s\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsection{Partiality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In HOL, all functions are total. A function \isa{f} applied to + \isa{x} always has the value \isa{f\ x}, and there is no notion + of undefinedness. + This is why we have to do termination + proofs when defining functions: The proof justifies that the + function can be defined by wellfounded recursion. + + However, the \cmd{function} package does support partiality to a + certain extent. Let's look at the following function which looks + for a zero of a given function f.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ {\isacharparenleft}if\ f\ n\ {\isacharequal}\ {\isadigit{0}}\ then\ n\ else\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Clearly, any attempt of a termination proof must fail. And without + that, we do not get the usual rules \isa{findzero{\isachardot}simps} and + \isa{findzero{\isachardot}induct}. So what was the definition good for at all?% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Domain predicates% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The trick is that Isabelle has not only defined the function \isa{findzero}, but also + a predicate \isa{findzero{\isacharunderscore}dom} that characterizes the values where the function + terminates: the \emph{domain} of the function. If we treat a + partial function just as a total function with an additional domain + predicate, we can derive simplification and + induction rules as we do for total functions. They are guarded + by domain conditions and are called \isa{psimps} and \isa{pinduct}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% +findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}\ {\isasymLongrightarrow}\isanewline +findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% +\end{isabelle}\end{minipage} + \hfill(\isa{findzero{\isachardot}psimps}) + \vspace{1em} + + \noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% +{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}{\isacharcomma}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}{\isacharparenright}{\isacharsemicolon}\isanewline +\isaindent{\ }{\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ n{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% +\end{isabelle}\end{minipage} + \hfill(\isa{findzero{\isachardot}pinduct})% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Remember that all we + are doing here is use some tricks to make a total function appear + as if it was partial. We can still write the term \isa{findzero\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isacharparenright}\ {\isadigit{0}}} and like any other term of type \isa{nat} it is equal + to some natural number, although we might not be able to find out + which one. The function is \emph{underdefined}. + + But it is defined enough to prove something interesting about it. We + can prove that if \isa{findzero\ f\ n} + terminates, it indeed returns a zero of \isa{f}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +\noindent We apply induction as usual, but using the partial induction + rule:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ f\ n\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent This gives the following subgoals: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isasymrbrakk}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ }{\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}% +\end{isabelle} + + \noindent The hypothesis in our lemma was used to satisfy the first premise in + the induction rule. However, we also get \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} as a local assumption in the induction step. This + allows to unfold \isa{findzero\ f\ n} using the \isa{psimps} + rule, and the rest is trivial. Since the \isa{psimps} rules carry the + \isa{{\isacharbrackleft}simp{\isacharbrackright}} attribute by default, we just need a single step:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ simp\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Proofs about partial functions are often not harder than for total + functions. Fig.~\ref{findzero_isar} shows a slightly more + complicated proof written in Isar. It is verbose enough to show how + partiality comes into play: From the partial induction, we get an + additional domain condition hypothesis. Observe how this condition + is applied when calls to \isa{findzero} are unfolded.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ f\ n\ \isacommand{assume}\isamarkupfalse% +\ dom{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ IH{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ x{\isacharunderscore}range{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ dom\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ x{\isacharunderscore}range\ \isacommand{show}\isamarkupfalse% +\ False\ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ x{\isacharunderscore}range\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ n\ {\isasymor}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\ \ \isacommand{thus}\isamarkupfalse% +\ {\isachardoublequoteopen}f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ dom\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ IH\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{A proof about a partial function}\label{findzero_isar} +\end{figure} +% +\isamarkupsubsection{Partial termination proofs% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Now that we have proved some interesting properties about our + function, we should turn to the domain predicate and see if it is + actually true for some values. Otherwise we would have just proved + lemmas with \isa{False} as a premise. + + Essentially, we need some introduction rules for \isa{findzero{\isacharunderscore}dom}. The function package can prove such domain + introduction rules automatically. But since they are not used very + often (they are almost never needed if the function is total), this + functionality is disabled by default for efficiency reasons. So we have to go + back and ask for them explicitly by passing the \isa{{\isacharparenleft}domintros{\isacharparenright}} option to the function package: + +\vspace{1ex} +\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% +\cmd{where}\isanewline% +\ \ \ldots\\ + + \noindent Now the package has proved an introduction rule for \isa{findzero{\isacharunderscore}dom}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ findzero{\isachardot}domintros% +\begin{isamarkuptext}% +\begin{isabelle}% +{\isacharparenleft}{\isadigit{0}}\ {\isacharless}\ {\isacharquery}f\ {\isacharquery}n\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% +\end{isabelle} + + Domain introduction rules allow to show that a given value lies in the + domain of a function, if the arguments of all recursive calls + are in the domain as well. They allow to do a \qt{single step} in a + termination proof. Usually, you want to combine them with a suitable + induction principle. + + Since our function increases its argument at recursive calls, we + need an induction principle which works \qt{backwards}. We will use + \isa{inc{\isacharunderscore}induct}, which allows to do induction from a fixed number + \qt{downwards}: + + \begin{center}\isa{{\isasymlbrakk}{\isacharquery}i\ {\isasymle}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharquery}j{\isacharsemicolon}\ {\isasymAnd}i{\isachardot}\ {\isasymlbrakk}i\ {\isacharless}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharparenleft}Suc\ i{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ i{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}i}\hfill(\isa{inc{\isacharunderscore}induct})\end{center} + + Figure \ref{findzero_term} gives a detailed Isar proof of the fact + that \isa{findzero} terminates if there is a zero which is greater + or equal to \isa{n}. First we derive two useful rules which will + solve the base case and the step case of the induction. The + induction is then straightforward, except for the unusual induction + principle.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}termination{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isasymge}\ n{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\ \isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ base{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ {\isacharparenleft}simp\ add{\isacharcolon}{\isacharbackquoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isacharbackquoteclose}{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ step{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}i{\isachardot}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}\ \isanewline +\ \ \ \ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ simp\isanewline +\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}x\ {\isasymge}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ base{\isacharparenright}\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ i\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{thus}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ step{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{Termination proof for \isa{findzero}}\label{findzero_term} +\end{figure} +% +\begin{isamarkuptext}% +Again, the proof given in Fig.~\ref{findzero_term} has a lot of + detail in order to explain the principles. Using more automation, we + can also have a short proof:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}termination{\isacharunderscore}short{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ zero{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isachargreater}{\isacharequal}\ n{\isachardoublequoteclose}\ \isanewline +\ \ \isakeyword{assumes}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ zero\isanewline +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\ {\isacharparenleft}auto\ intro{\isacharcolon}\ findzero{\isachardot}domintros{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent It is simple to combine the partial correctness result with the + termination lemma:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}total{\isacharunderscore}correctness{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}blast\ intro{\isacharcolon}\ findzero{\isacharunderscore}zero\ findzero{\isacharunderscore}termination{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Definition of the domain predicate% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Sometimes it is useful to know what the definition of the domain + predicate looks like. Actually, \isa{findzero{\isacharunderscore}dom} is just an + abbreviation: + + \begin{isabelle}% +findzero{\isacharunderscore}dom\ {\isasymequiv}\ accp\ findzero{\isacharunderscore}rel% +\end{isabelle} + + The domain predicate is the \emph{accessible part} of a relation \isa{findzero{\isacharunderscore}rel}, which was also created internally by the function + package. \isa{findzero{\isacharunderscore}rel} is just a normal + inductive predicate, so we can inspect its definition by + looking at the introduction rules \isa{findzero{\isacharunderscore}rel{\isachardot}intros}. + In our case there is just a single rule: + + \begin{isabelle}% +{\isacharquery}f\ {\isacharquery}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}rel\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% +\end{isabelle} + + The predicate \isa{findzero{\isacharunderscore}rel} + describes the \emph{recursion relation} of the function + definition. The recursion relation is a binary relation on + the arguments of the function that relates each argument to its + recursive calls. In general, there is one introduction rule for each + recursive call. + + The predicate \isa{findzero{\isacharunderscore}dom} is the accessible part of + that relation. An argument belongs to the accessible part, if it can + be reached in a finite number of steps (cf.~its definition in \isa{Wellfounded{\isachardot}thy}). + + Since the domain predicate is just an abbreviation, you can use + lemmas for \isa{accp} and \isa{findzero{\isacharunderscore}rel} directly. Some + lemmas which are occasionally useful are \isa{accpI}, \isa{accp{\isacharunderscore}downward}, and of course the introduction and elimination rules + for the recursion relation \isa{findzero{\isachardot}intros} and \isa{findzero{\isachardot}cases}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{A Useful Special Case: Tail recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The domain predicate is our trick that allows us to model partiality + in a world of total functions. The downside of this is that we have + to carry it around all the time. The termination proof above allowed + us to replace the abstract \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} by the more + concrete \isa{n\ {\isasymle}\ x\ {\isasymand}\ f\ x\ {\isacharequal}\ {\isadigit{0}}}, but the condition is still + there and can only be discharged for special cases. + In particular, the domain predicate guards the unfolding of our + function, since it is there as a condition in the \isa{psimp} + rules. + + Now there is an important special case: We can actually get rid + of the condition in the simplification rules, \emph{if the function + is tail-recursive}. The reason is that for all tail-recursive + equations there is a total function satisfying them, even if they + are non-terminating. + +% A function is tail recursive, if each call to the function is either +% equal +% +% So the outer form of the +% +%if it can be written in the following +% form: +% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} + + + The function package internally does the right construction and can + derive the unconditional simp rules, if we ask it to do so. Luckily, + our \isa{findzero} function is tail-recursive, so we can just go + back and add another option to the \cmd{function} command: + +\vspace{1ex} +\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharcomma}\ tailrec{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% +\cmd{where}\isanewline% +\ \ \ldots\\% + + + \noindent Now, we actually get unconditional simplification rules, even + though the function is partial:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ findzero{\isachardot}simps% +\begin{isamarkuptext}% +\begin{isabelle}% +findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% +\end{isabelle} + + \noindent Of course these would make the simplifier loop, so we better remove + them from the simpset:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{declare}\isamarkupfalse% +\ findzero{\isachardot}simps{\isacharbrackleft}simp\ del{\isacharbrackright}% +\begin{isamarkuptext}% +Getting rid of the domain conditions in the simplification rules is + not only useful because it simplifies proofs. It is also required in + order to use Isabelle's code generator to generate ML code + from a function definition. + Since the code generator only works with equations, it cannot be + used with \isa{psimp} rules. Thus, in order to generate code for + partial functions, they must be defined as a tail recursion. + Luckily, many functions have a relatively natural tail recursive + definition.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Nested recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Recursive calls which are nested in one another frequently cause + complications, since their termination proof can depend on a partial + correctness property of the function itself. + + As a small example, we define the \qt{nested zero} function:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ nz\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}nz\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}nz\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ nz\ {\isacharparenleft}nz\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +If we attempt to prove termination using the identity measure on + naturals, this fails:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ auto% +\begin{isamarkuptxt}% +We get stuck with the subgoal + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharless}\ Suc\ n% +\end{isabelle} + + Of course this statement is true, since we know that \isa{nz} is + the zero function. And in fact we have no problem proving this + property by induction.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}nz{\isachardot}pinduct{\isacharparenright}\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +We formulate this as a partial correctness lemma with the condition + \isa{nz{\isacharunderscore}dom\ n}. This allows us to prove it with the \isa{pinduct} rule before we have proved termination. With this lemma, + the termination proof works as expected:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ {\isacharparenleft}auto\ simp{\isacharcolon}\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +As a general strategy, one should prove the statements needed for + termination as a partial property first. Then they can be used to do + the termination proof. This also works for less trivial + examples. Figure \ref{f91} defines the 91-function, a well-known + challenge problem due to John McCarthy, and proves its termination.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +\isacommand{function}\isamarkupfalse% +\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n\ then\ n\ {\isacharminus}\ {\isadigit{1}}{\isadigit{0}}\ else\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate{\isacharcolon}\ \isanewline +\ \ \isakeyword{assumes}\ trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ n{\isachardoublequoteclose}\ \isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}n\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ trm\ \isacommand{by}\isamarkupfalse% +\ induct\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{termination}\isamarkupfalse% +\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{let}\isamarkupfalse% +\ {\isacharquery}R\ {\isacharequal}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{1}}\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}wf\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isachardoublequoteclose}\ % +\isamarkupcmt{Assumptions for both calls% +} +\isanewline +\isanewline +\ \ \isacommand{thus}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\ % +\isamarkupcmt{Inner call% +} +\isanewline +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ inner{\isacharunderscore}trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ % +\isamarkupcmt{Outer call% +} +\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupfalse\isabellestyle{tt} +\end{minipage} +\vspace{6pt}\hrule +\caption{McCarthy's 91-function}\label{f91} +\end{figure} +% +\isamarkupsection{Higher-Order Recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Higher-order recursion occurs when recursive calls + are passed as arguments to higher-order combinators such as \isa{map}, \isa{filter} etc. + As an example, imagine a datatype of n-ary trees:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isacharprime}a\ tree\ {\isacharequal}\ \isanewline +\ \ Leaf\ {\isacharprime}a\ \isanewline +{\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ list{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We can define a function which swaps the left and right subtrees recursively, using the + list functions \isa{rev} and \isa{map}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ mirror\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Leaf\ n{\isacharparenright}\ {\isacharequal}\ Leaf\ n{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Branch\ l{\isacharparenright}\ {\isacharequal}\ Branch\ {\isacharparenleft}rev\ {\isacharparenleft}map\ mirror\ l{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Although the definition is accepted without problems, let us look at the termination proof:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +% +\begin{isamarkuptxt}% +As usual, we have to give a wellfounded relation, such that the + arguments of the recursive calls get smaller. But what exactly are + the arguments of the recursive calls when mirror is given as an + argument to \isa{map}? Isabelle gives us the + subgoals + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharquery}R\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}l\ x{\isachardot}\ x\ {\isasymin}\ set\ l\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ Branch\ l{\isacharparenright}\ {\isasymin}\ {\isacharquery}R% +\end{isabelle} + + So the system seems to know that \isa{map} only + applies the recursive call \isa{mirror} to elements + of \isa{l}, which is essential for the termination proof. + + This knowledge about \isa{map} is encoded in so-called congruence rules, + which are special theorems known to the \cmd{function} command. The + rule for \isa{map} is + + \begin{isabelle}% +{\isasymlbrakk}{\isacharquery}xs\ {\isacharequal}\ {\isacharquery}ys{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ set\ {\isacharquery}ys\ {\isasymLongrightarrow}\ {\isacharquery}f\ x\ {\isacharequal}\ {\isacharquery}g\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ map\ {\isacharquery}f\ {\isacharquery}xs\ {\isacharequal}\ map\ {\isacharquery}g\ {\isacharquery}ys% +\end{isabelle} + + You can read this in the following way: Two applications of \isa{map} are equal, if the list arguments are equal and the functions + coincide on the elements of the list. This means that for the value + \isa{map\ f\ l} we only have to know how \isa{f} behaves on + the elements of \isa{l}. + + Usually, one such congruence rule is + needed for each higher-order construct that is used when defining + new functions. In fact, even basic functions like \isa{If} and \isa{Let} are handled by this mechanism. The congruence + rule for \isa{If} states that the \isa{then} branch is only + relevant if the condition is true, and the \isa{else} branch only if it + is false: + + \begin{isabelle}% +{\isasymlbrakk}{\isacharquery}b\ {\isacharequal}\ {\isacharquery}c{\isacharsemicolon}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}u{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}v{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharparenleft}if\ {\isacharquery}b\ then\ {\isacharquery}x\ else\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}c\ then\ {\isacharquery}u\ else\ {\isacharquery}v{\isacharparenright}% +\end{isabelle} + + Congruence rules can be added to the + function package by giving them the \isa{fundef{\isacharunderscore}cong} attribute. + + The constructs that are predefined in Isabelle, usually + come with the respective congruence rules. + But if you define your own higher-order functions, you may have to + state and prove the required congruence rules yourself, if you want to use your + functions in recursive definitions.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Congruence Rules and Evaluation Order% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Higher order logic differs from functional programming languages in + that it has no built-in notion of evaluation order. A program is + just a set of equations, and it is not specified how they must be + evaluated. + + However for the purpose of function definition, we must talk about + evaluation order implicitly, when we reason about termination. + Congruence rules express that a certain evaluation order is + consistent with the logical definition. + + Consider the following function.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{0}}\ {\isasymor}\ f\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +For this definition, the termination proof fails. The default configuration + specifies no congruence rule for disjunction. We have to add a + congruence rule that specifies left-to-right evaluation order: + + \vspace{1ex} + \noindent \isa{{\isasymlbrakk}{\isacharquery}P\ {\isacharequal}\ {\isacharquery}P{\isacharprime}{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}P{\isacharprime}\ {\isasymLongrightarrow}\ {\isacharquery}Q\ {\isacharequal}\ {\isacharquery}Q{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharquery}P\ {\isasymor}\ {\isacharquery}Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}P{\isacharprime}\ {\isasymor}\ {\isacharquery}Q{\isacharprime}{\isacharparenright}}\hfill(\isa{disj{\isacharunderscore}cong}) + \vspace{1ex} + + Now the definition works without problems. Note how the termination + proof depends on the extra condition that we get from the congruence + rule. + + However, as evaluation is not a hard-wired concept, we + could just turn everything around by declaring a different + congruence rule. Then we can make the reverse definition:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ disj{\isacharunderscore}cong{\isadigit{2}}{\isacharbrackleft}fundef{\isacharunderscore}cong{\isacharbrackright}{\isacharcolon}\ \isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isasymnot}\ Q{\isacharprime}\ {\isasymLongrightarrow}\ P\ {\isacharequal}\ P{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}Q\ {\isacharequal}\ Q{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}P{\isacharprime}\ {\isasymor}\ Q{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{fun}\isamarkupfalse% +\ f{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}f{\isacharprime}\ n\ {\isacharequal}\ {\isacharparenleft}f{\isacharprime}\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymor}\ n\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent These examples show that, in general, there is no \qt{best} set of + congruence rules. + + However, such tweaking should rarely be necessary in + practice, as most of the time, the default set of congruence rules + works well.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/Thy/document/session.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/document/session.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,6 @@ +\input{Functions.tex} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/conclusion.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/conclusion.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,7 @@ +\section{Conclusion} + +\fixme{} + + + + diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/functions.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/functions.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,91 @@ + +\documentclass[a4paper,fleqn]{article} + +\usepackage{latexsym,graphicx} +\usepackage[refpage]{nomencl} +\usepackage{../iman,../extra,../isar,../proof} +\usepackage{../isabelle,../isabellesym} +\usepackage{style} +\usepackage{mathpartir} +\usepackage{amsthm} +\usepackage{../pdfsetup} + +\newcommand{\cmd}[1]{\isacommand{#1}} + +\newcommand{\isasymINFIX}{\cmd{infix}} +\newcommand{\isasymLOCALE}{\cmd{locale}} +\newcommand{\isasymINCLUDES}{\cmd{includes}} +\newcommand{\isasymDATATYPE}{\cmd{datatype}} +\newcommand{\isasymAXCLASS}{\cmd{axclass}} +\newcommand{\isasymDEFINES}{\cmd{defines}} +\newcommand{\isasymNOTES}{\cmd{notes}} +\newcommand{\isasymCLASS}{\cmd{class}} +\newcommand{\isasymINSTANCE}{\cmd{instance}} +\newcommand{\isasymLEMMA}{\cmd{lemma}} +\newcommand{\isasymPROOF}{\cmd{proof}} +\newcommand{\isasymQED}{\cmd{qed}} +\newcommand{\isasymFIX}{\cmd{fix}} +\newcommand{\isasymASSUME}{\cmd{assume}} +\newcommand{\isasymSHOW}{\cmd{show}} +\newcommand{\isasymNOTE}{\cmd{note}} +\newcommand{\isasymCODEGEN}{\cmd{code\_gen}} +\newcommand{\isasymPRINTCODETHMS}{\cmd{print\_codethms}} +\newcommand{\isasymFUN}{\cmd{fun}} +\newcommand{\isasymFUNCTION}{\cmd{function}} +\newcommand{\isasymPRIMREC}{\cmd{primrec}} +\newcommand{\isasymRECDEF}{\cmd{recdef}} + +\newcommand{\qt}[1]{``#1''} +\newcommand{\qtt}[1]{"{}{#1}"{}} +\newcommand{\qn}[1]{\emph{#1}} +\newcommand{\strong}[1]{{\bfseries #1}} +\newcommand{\fixme}[1][!]{\strong{FIXME: #1}} + +\newtheorem{exercise}{Exercise}{\bf}{\itshape} +%\newtheorem*{thmstar}{Theorem}{\bf}{\itshape} + +\hyphenation{Isabelle} +\hyphenation{Isar} + +\isadroptag{theory} +\title{Defining Recursive Functions in Isabelle/HOL} +\author{Alexander Krauss} + +\isabellestyle{tt} +\renewcommand{\isastyletxt}{\isastyletext}% use same formatting for txt and text + +\begin{document} + +\date{\ \\} +\maketitle + +\begin{abstract} + This tutorial describes the use of the new \emph{function} package, + which provides general recursive function definitions for Isabelle/HOL. + We start with very simple examples and then gradually move on to more + advanced topics such as manual termination proofs, nested recursion, + partiality, tail recursion and congruence rules. +\end{abstract} + +%\thispagestyle{empty}\clearpage + +%\pagenumbering{roman} +%\clearfirst + +\input{intro.tex} +\input{Thy/document/Functions.tex} +%\input{conclusion.tex} + +\begingroup +%\tocentry{\bibname} +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{../manual} +\endgroup + +\end{document} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/intro.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/intro.tex Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,55 @@ +\section{Introduction} + +Starting from Isabelle 2007, new facilities for recursive +function definitions~\cite{krauss2006} are available. They provide +better support for general recursive definitions than previous +packages. But despite all tool support, function definitions can +sometimes be a difficult thing. + +This tutorial is an example-guided introduction to the practical use +of the package and related tools. It should help you get started with +defining functions quickly. For the more difficult definitions we will +discuss what problems can arise, and how they can be solved. + +We assume that you have mastered the fundamentals of Isabelle/HOL +and are able to write basic specifications and proofs. To start out +with Isabelle in general, consult the Isabelle/HOL tutorial +\cite{isa-tutorial}. + + + +\paragraph{Structure of this tutorial.} +Section 2 introduces the syntax and basic operation of the \cmd{fun} +command, which provides full automation with reasonable default +behavior. The impatient reader can stop after that +section, and consult the remaining sections only when needed. +Section 3 introduces the more verbose \cmd{function} command which +gives fine-grained control. This form should be used +whenever the short form fails. +After that we discuss more specialized issues: +termination, mutual, nested and higher-order recursion, partiality, pattern matching +and others. + + +\paragraph{Some background.} +Following the LCF tradition, the package is realized as a definitional +extension: Recursive definitions are internally transformed into a +non-recursive form, such that the function can be defined using +standard definition facilities. Then the recursive specification is +derived from the primitive definition. This is a complex task, but it +is fully automated and mostly transparent to the user. Definitional +extensions are valuable because they are conservative by construction: +The \qt{new} concept of general wellfounded recursion is completely reduced +to existing principles. + + + + +The new \cmd{function} command, and its short form \cmd{fun} have mostly +replaced the traditional \cmd{recdef} command \cite{slind-tfl}. They solve +a few of technical issues around \cmd{recdef}, and allow definitions +which were not previously possible. + + + + diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/isabelle_isar.eps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/isabelle_isar.eps Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,6488 @@ +%!PS-Adobe-2.0 EPSF-1.2 +%%Title: isabelle_any +%%Creator: FreeHand 5.5 +%%CreationDate: 24.09.1998 21:04 Uhr +%%BoundingBox: 0 0 202 178 +%%FHPathName:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any +%ALDOriginalFile:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any +%ALDBoundingBox: -153 -386 442 456 +%%FHPageNum:1 +%%DocumentSuppliedResources: procset Altsys_header 4 0 +%%ColorUsage: Color +%%DocumentProcessColors: Cyan Magenta Yellow Black +%%DocumentNeededResources: font Symbol +%%+ font ZapfHumanist601BT-Bold +%%DocumentFonts: Symbol +%%+ ZapfHumanist601BT-Bold +%%DocumentNeededFonts: Symbol +%%+ ZapfHumanist601BT-Bold +%%EndComments +%!PS-AdobeFont-1.0: ZapfHumanist601BT-Bold 003.001 +%%CreationDate: Mon Jun 22 16:09:28 1992 +%%VMusage: 35200 38400 +% Bitstream Type 1 Font Program +% Copyright 1990-1992 as an unpublished work by Bitstream Inc., Cambridge, MA. +% All rights reserved. +% Confidential and proprietary to Bitstream Inc. +% U.S. GOVERNMENT RESTRICTED RIGHTS +% This software typeface product is provided with RESTRICTED RIGHTS. Use, +% duplication or disclosure by the Government is subject to restrictions +% as set forth in the license agreement and in FAR 52.227-19 (c) (2) (May, 1987), +% when applicable, or the applicable provisions of the DOD FAR supplement +% 252.227-7013 subdivision (a) (15) (April, 1988) or subdivision (a) (17) +% (April, 1988). Contractor/manufacturer is Bitstream Inc., +% 215 First Street, Cambridge, MA 02142. +% Bitstream is a registered trademark of Bitstream Inc. +11 dict begin +/FontInfo 9 dict dup begin + /version (003.001) readonly def + /Notice (Copyright 1990-1992 as an unpublished work by Bitstream Inc. All rights reserved. Confidential.) readonly def + /FullName (Zapf Humanist 601 Bold) readonly def + /FamilyName (Zapf Humanist 601) readonly def + /Weight (Bold) readonly def + /ItalicAngle 0 def + /isFixedPitch false def + /UnderlinePosition -136 def + /UnderlineThickness 85 def +end readonly def +/FontName /ZapfHumanist601BT-Bold def +/PaintType 0 def +/FontType 1 def +/FontMatrix [0.001 0 0 0.001 0 0] readonly def +/Encoding StandardEncoding def +/FontBBox {-167 -275 1170 962} readonly def +/UniqueID 15530396 def +currentdict end +currentfile eexec +a2951840838a4133839ca9d22e2b99f2b61c767cd675080aacfcb24e19cd +1336739bb64994c56737090b4cec92c9945ff0745ef7ffc61bb0a9a3b849 +e7e98740e56c0b5af787559cc6956ab31e33cf8553d55c0b0e818ef5ec6b +f48162eac42e7380ca921dae1c82b38fd6bcf2001abb5d001a56157094cf +e27d8f4eac9693e88372d20358b47e0c3876558ebf757a1fbc5c1cddf62b +3c57bf727ef1c4879422c142a084d1c7462ac293e097fabe3a3ecfcd8271 +f259833bac7912707218ec9a3063bf7385e02d8c1058ac06df00b33b8c01 +8768b278010eb4dd58c7ba59321899741cb7215d8a55bee8d3398c887f02 +e1f4869387f89141de693fcb429c7884c22dcdeddcaa62b7f5060249dfab +cfc351201f7d188b6ed68a228abda4d33b3d269ac09cde172bc045e67449 +c0f25d224efbe8c9f9d2968a01edbfb039123c365ed0db58ad38aabe015b +8881191dd23092f6d53d5c1cd68ebd038e098d32cb24b433b7d5d89c28ee +05ea0b6070bb785a2974b5a160ee4cf8b6d8c73445d36720af0530441cd9 +61bc0c367f1af1ec1c5ab7255ddda153c1868aba58cd5b44835535d85326 +5d7fed5ff7118adb5d5b76cc3b72e5ff27e21eb857261b3afb7688fca12d +1663b0d8bdc1dd47a84b65b47d3e76d3b8fa8b319f17e1bb22b45a7482fd +f9ad1b6129e09ae47f15cd2447484cd2d64f59ab0f2f876c81e7d87ccdf9 +005aa8fc093d02db51a075d571e925f2d309a1c535a1e59d34215c6cd33e +3c38997b4956461f376399901a8d0943dca6a335baac93fc8482c0659f04 +329c6f040b35828ea6dd1bd1858f2a9be4ef77731b5b75a1c536c6bc9479 +0821e5d88f6e2981835dbfd65ec254ebcf2cf49c917d121cd3bbb476a12b +69c15f17d9c17bb15ad1e7d31d2afcf58c8f0ad526d68615a0f1ac3b1d1c +d3beafeea3cf56c8f8a66367c70df9159f0b1b3157ccfd010045c4718e0e +625c0891e85790c9b97b85517c74c9d55eaca31a01cddc64898bf0eeadf3 +53391a185e507dcb0a6f52661a56357ac818dfc740a540aadf02f4e7a79d +8008a77cd30abde337025b01217d6a68c306abe145b7260c3478fa5f366f +b2d37259ead8a8ec2db2f09ae0eb3a682d27b0d73a60935f80254c24426a +003a87a29a4370cbf1b2ef1e19ad8466ec725fd5b463d06011a5e0da2258 +ff6c1483c4532bc21f2ed9b99b929b2800ddefc1a98d12ba085adc210bac +e0274b69e24d16af015a51ca73edf779a7534a887aa780337ad966839881 +edc22ba72038aa1a96a6deba24ad676795da711b92f8cf5f54cb4322ec04 +40ef9e15b11e3005f3ff69376ecb29bb66e8fc1b685f2b05fb162fcb35aa +d7eb2a8ec39b97ab1ff05ef02f8dbbc12085a5cd252cc4010fab7f63dfd5 +7fa1be86f724d37db5faef17ae8e8e537444e8e9db917b270344183473af +7f47d5703a160d8ef1e772438620d3336b2fbcf6433612e4b5e64fae0329 +8a3c4010c17d93f13ba66d549c69dd58c7d26ddc90285fed831918563a16 +2a7ac2511e2f18c9eb3df905a9dcba65a31cc1c39f07458abb11b4c60346 +aea19070e126982f1dde336e79be0ecd69a8afbe2493d064d4d2ff38788b +b3038125961302db9761403c3b8019ec641e107425002205a77ae2ae0f4f +7550d623dd03f0ec0199f42a9a8b89514e2e21baca9b3c8c96ca48cbf9dc +ee6d6241d713e014b49e83ad85e62a6b2f70b58e4cc72d41ea6fcbdd3b5c +890c8af0d24200658773b1628c6cc9aaaabb08865ee4c7ff2c513ad7aa23 +155a321213fa94731683a3e282a0e60aa3c87aade3da231465bdd9097f2c +89a1af8e5b9649143f2d9482546501ea97e8bea2f5d5eea97d4f19bb6835 +3138d3babb2461c08d537491aaede1f23d734c6f099eb5bef6e2ffaaf138 +e5ab71b8b41599091037e440127a4eaedf208c20c8a2fc62eadab191d1ab +4d5531f826aa6b9fff2797a7f54673e0a3fae09a93a0dfafb8b11d60dc69 +5acf9b7e1a47c31d0b5a0b85b7b50cddff5ac831651d9c7469c2139c7a89 +7d2f868f36c65156921803eccfdbdd1618595ab6d2a9230ef523a1b5ee51 +f2a0d200fc0e94aff7f546593ff2a3eb865d129895af01b8ab6e4616fe20 +9123b6e2b7e0817adc3cdb78ae8b0b1d75f2986ebd8fb24c4de92ac9e8c3 +6afa520636bcad2e6a03d11c268d97fa578561f6e7523e042c4cc73a0eac +7a841907450e83d8e7a8de4db5085f6e8b25dc85b59e018380f4b9523a7f +02cbeec989f0221b7681ec427519062b429dcd8fc2e4f81173519f88e2e4 +3798b90a84060920f6ae789afd6a9182e7fec87cd2d4235d37a65c3f3bcc +c742c89cbe5f3e2ba6c4f40ebba162e12569d20684cc167685285a087e7a +0a995fe1939bf25c09553512ba2cf77ef21d2ef8da1c73ba6e5826f4099e +27d8bc7b3545fc592b75228e70127c15a382774357457cd4586d80dc0bd6 +065aee32acfd5c0523303cece85a3dbf46853b917618c0330146f527c15b +dbb9f6526964368b2b8593eed1551dad75659565d54c6a0a52da7a8e366f +dd009ef853491c0fb995e19933cba1dbdc8902721c3ea6017ffdd5851cb8 +3c8bada46075ac121afe13a70e87529e40693157adcc999ed4657e017adf +f7dbac4bc0d204f204c6f47b769aaf714f9ec1d25226f24d0a1b53e28ac5 +374ab99755852c1431b2486df5fd637e2005a25303345a1c95a15a1189ba +f6f6883de1ad46d48427b137c2003d210ab2b2f5680f2633939f289d7553 +eb943adf8127f1c3ee7d6453b5566393700ad74ab86eb9a89f8b0380af55 +6b62f51b7dbd0c5dcc9a9fb658944d7ad5845d58dedc2d38200d0ef7cb0f +76041dc104ef3ab89c1dc2f6a75635d48051c8a7dd9f5e60253a53957ec8 +9d1266566d7ed20d79dfc2807b397d7cf056bdaccdb72528a88aa4987682 +c909b2fe1e35a71c2f29e89a2bf32173967e79610367ce4574ba6a1cc031 +cfb176fc0313f74f91a866ef9954b95b29caf917a6b919586f54d23cb7ce +23305886ae7760ebd6263df0d3c511ac7afc361df78bc2621f66d3268b99 +078fa59124f0eb9476496c938eb4584e87455dc6f2faa999e938460b31c6 +28021c652acfa12d4556aa4302bbcd043e60389480b796c3fc0b2e51b81e +c2afa4a34335318a1c5a842dcaa120df414acba2e79ab5cc61c99e98108c +5cb907a96b30d731131782f9df79aabfc16a20ace8852d047497982e11c8 +26321addf679de8a96a2d18743f6f2c3b2bc397370b10ad273fcfb76b48b +9dad27cf89ca2c370149cd48ab956e9bbce01cbf8e1f0c661b99cf19b36e +87b6165dd85ae3f3674525e17d85971093e110520d17f2d6253611e35ec9 +e17000e48e2926454c1e8eb4098e0115b49536f2c33505eb5e574f7a414b +e176398c5ddf6d846ea5ddf2a5e94c0422e0115c57a8c9e56bf8ba962c82 +698c96bd6138baaca7347e44352cc62df4eeba364954ad921a5a43280980 +264df4a7fb29d005423179f7bd1d98b4280d62ce23c927551f1ffc2b8f17 +0a9c23656c0c91b640cdcfdbd88089ffb28d3ac68bad25dbbed82c083696 +1f9f86a6183cc1061ffdb32279796569d05b31c946955402d0be1fb9f2bf +304d1ad8e1e357be49e6e2ee67f7f1e7bc699d46a5f7853fe659ba2e1930 +0d3e3ea658b9862701dcab08fdd23bf1d751777f25efbe9e02d12b5612b3 +c3fc2275190346b94ec4024e4ade08e54d75c0b4c48f4956b9182e8ce997 +74b54da4a9318c099d89f1ce3b6803a95f48b9fb8b845372be25e54478e8 +49e4707ea03a36e134efa661e4e6250d89649ae074cfd9d6b9e2071c1099 +3b8a5a5ebc3e1cb228c97565aef7f254e3f90af8a3dd281c83792755719d +c6a5b3bab4aa6be5afe9624050eee8dfb13b018f4088c932cd48ace38dfe +b1b4218dba8f7fada6628076acf1b54db0c95d4fb12232f1fa9d1ba848f9 +fe80c65b75d6946c00fe78839142c5302707d9269b24565dbcc551aca060 +b4d0f99b961dd3cc795a982063ac42e9fc81fc98add42744a9f92e74b00d +637ee4606ea2099b6c763493c3159f8e52a90dafca682010c0e92bc9038a +10abb066c75c8d97f7ad6fb1a37136e52cf2093c4fa485fe12adad10e4d0 +83b78b023628ddc5326cbf8392516027a9f3de4945f93488e4a1997efd2a +22c2c136dbac1bdb829e082beac48cdd06dcb17bacf09451c7b636bd49a8 +fc60cb1d2292250bea78e1dd276725ab4c526b66ddabf4e1b2bf0a2571df +2490df70735f5da321fac74fe4fab54444e53dace9832cff38a70c58104a +4f0c0545dcf7a1a9ecb54e0e32d6d8195d30b2c98a567741fcf325a4ddeb +244a1a35676e246ddc835fac13b569f35f22ec93191eca3efbe17ff9a950 +d08f863710b2bbecec969068c498eb2338b68f3fc3f5447449fe4de2b065 +e068ecd9255d369b2bb6c4b0b7423774bed69294758aca2bdb1a8d5bf618 +d3fa09462f7838e8a79b7a53bebe6dacb0a1561eaa074bc6f069f6a06fb2 +c4a5cb13e2172bce9be659a82665da5cded73da84322bb16aa6e19ac1958 +7515cb5d2b65e65e803f76700ce5efd3df7fe4ed431fae0e8e286b1d5056 +a0d18df926b2be7a93c503ab903abd4788680a6201fdc299f2cb5d6a9b6e +2048109c8d1fb633a54128938594b2cce86a7e0185e7d59e6536584039ec +9e30ff7be6ddba9fdba82de7415fdc47de84d97afb1aa3ba495bd91dee9d +f3b21ee1164987dd8510925087cd0919f1085cba6e4dd3c7384d26667f94 +ad7f736a60d8bd76dfaa4b53c23536fc309ff2317c48ee0107ff2ca4d1b3 +f78c5a27b901c931128bdb636094ef0cd543a5b62b7dbe10ed83aed5780c +a16067a4a7e8b7c5bf8a8e822149bc1b5bcdabe13a7f6aa6eaeff24a42f4 +a58a2b70f545103540169137fda9abb589f734b6776cb50402d6123ce802 +10dce05e3697a98c9411cf60a02c278c91e03d540b936cd00c668960e357 +1aeaf4d94cfb496b259ec0d8fdba9199fb46634ff177bc8d310ea1314eef +d46c927a981c58e88743ed4e07d80fe841edee812e3053412bf2e710146c +b25dec8ea70c38bb1f6e4db3c2e8ba521963c1584eeb60ea1e9555058f13 +e98307c13cbd15c26b611f543149b1ddf88dd6296ae703f58daeb67f1b03 +ab5b24c72d5770cb9d8ed242c4faaad1dd940ada00e98ff3a61799d13355 +aba916910aa9a6e5ee8af438d0ba8235346fcd139b9d2cb7db7bd3f298a3 +94ff0aff3b9042f32a004e042c346a5ea35917f229848a9c5a32909b0090 +4aa715640277a6ada99f8b2927fda22899ff1347f42bac73e2bd4bbf3945 +55fd7dd30d5c6dadf5c259fdb2455d8c607de1c5da588e20159f18e4e8da +b714e532d888a0386c52c2b85964251af003ac9d10c0c8b9b3465e1dde48 +2e74a29e17a7cf6c9a43b5af1646f0b8508f98e9a1279ec3908073d89dcb +aa093e8dd1004c1ecccce0803095b0069d4be7a1eb14b02efc37d137dfe3 +f0190bc9628069abc257f45d0e050e60c7f5281277937dd986fcd5b94a2b +845a1a75addd74a142800f18ef408c08a2c2ad16a93298f148c0ae7d2990 +ded147f92f053809a60d4f992a227167aad5b5eb2bbe8a4a77dc77a08b72 +6acb34422e2532eec7e274e4a42d71ee584f5f3df5a3a5b379974ede73ab +5f1b0a2dbfcc8cfac4747ca26eb6030dc2f85a104b754432977850c093b9 +97ed90af711b544ff682e7b1eac82b2b7b44014b09c17ecf084c994a095d +9eeef5391305caf093b62ac9916f982a436d521fcf4d75c5b8e4d92266fd +e99a58aa39d7693ecd1796b8851761d64bbca39a6d5a0b4533ae47123327 +f98d0ad0e8b36625cc3647b55459552906d8a1d5766845ffac101980efcf +79657e365510be5db557cefef21193ca3cf3dad175ee2e7ae91d174fdc06 +2ff5c51ffe2f021122e96df042019d3a1883e662537ec1b69c11fbb6e750 +0132eabf5803c816153ecbff60ca3b3b39708c26cb1751afb2e65d8e5f4a +c4397a88fb1f112672fcdd24e4ba545c5b2a7968c17b62f8e2530a8acbff +cfca82c64b7abcab84e2c4a0a7ced67b15669301fe9ff2c756e70ff7ce33 +497be6acc4ac5617e1f043bd8a87416299a39bf17fc31c02d72d75fdc2a1 +e60669fa4d5e4a49d9afea2f53f4626680e9c0dfca223529efa415c4343a +b6067aa800c484457ea050eaaa5d3fafeedd0eec72f327e02c6b3912b5a8 +c404de4839c9c4a99da42681cde43316606a34c7d2f02269de1aab776857 +e668f35946af4d618d36d444bdc02b1f63ea25b6260b4fb606ac8575b5c9 +782a5de4037350d5753b1537537ccb008c454eeb264e6cd4727c999e240e +0ac89e95a896b67d54910a3531345f64198ad394b5ceb52881f1dd9e6beb +95862dc188d45b3e46aacb5fe40097947dab9bc3c1ee46bfc9b1b3ed6167 +efd0d65ceb043d7b24c1456676e4baa47b1209a315f199bb3a91f4374cd9 +cc0b40d3f09f19f8dd8a46915eee55eeeeb3c7b8f437106ee491ef0f4ff9 +2c5c6f779e0fbe7bd5293964bb645ca362b106abeb773571d9ae83b786a3 +d5a4ea3ea970daadc46cc5e6037f76fd20e0fffc47cf4e7af9522b91f96b +3297720fd45d9bc2200622ad2ca9445556c8a8202b1991bc63da360d021d +55be2528e043f803e08da99b91ab9cfc5e65b2655d78206b4aecd445a7b0 +1caa0d06b0a55e8f04b70b60b04a860c8e1ab439f4910051e3f7441b47c7 +8aa3ab8519f181a9e833f3242fa58d02ed76bf0031f71f9def8484ecced8 +b6e41aca56176b6b32a2443d12492c8a0f5ba8a3e227219dfd1dd23fcb48 +fcfd255dbbf3e198874e607399db8d8498e719f00e9ed8bdd96c88817606 +357a0063c23854e64ad4e59ddd5060845b2c4cddd00c40081458f8ee02c7 +303c11747bd104440046bf2d09794fca2c4beb23ed1b66d9ccb9a4dd57ad +a24943461ecc00704c916bdc621bfddb17913dfb0f3513b65f3ab015786a +caa51ee9546bc8ddf87e2e104137e35ddf8f8d23724e9a53824169bc7cfa +99562656e6f1c888d4dbff0b269c5d1e733e5f212d91297610201eb43249 +35e336dd0052738db2d64f3e89429903bb5c1810009cf766e9a06223dd2f +219b706394a121dc029af55c6ada9052af59682ef7c51e121cf16f0319ac +0aa9512ef900c548d673fe361da19052808797e958209072e145d46ec8cc +a89fafd76630eff30ae979973bdf0f8c9e469d8edd3b1c93731c72f976b7 +d81142bc15c376403f967affaa5f482efd57c6f91970729d16db851f0ed3 +ea7d82f409307b5b436886c1beda94a1fa3ab1b60686f6574c844fb2c0b3 +a07174dc4f27b4fed2f8bd4d5436be4b343e5efdf0400d235bd910255341 +a20770804a26f8437e9bce6da8e9f8258a343c7aee291f1510be306ae67a +ab1d7696453530c02fd153bbe49dbf62baad6146029cbd1656cdb76c952c +b93edfee76fe33832930be59636bb947e8ad285f20f663cccf484fba97d6 +7446c7b6c6f5857428bb1737d9ae801df75d9cb4d7bd59ef7a4cbadde928 +38f15d232005585d2e40483d2d3e08cc8f398bb43afedb84343c3ba3835d +0ba82a86dce859cf655f85e63e41365e0dbefcf511b9a27a2b6e66b2ad3a +c657902842287a317e46ceaa93b5088f09d53a65815b44538af90ad3b06b +4e5e2dc509f02e30a01e05201c67d4d39582bbe64e20b669f5fd787909a3 +30fc50a95b31426bbb57a4fbf9feacdc31f98bcf50da7e50c2bfc169c6fd +ccf213cdb878653bcea372968ea6e31fd30dd55434cc91c0af22179ce669 +a05493f195e12432c6173ae2ac3c94fb83f38210014a9f969ea2b44e99f5 +e5a7317e848d429ad62167a4fc5001149676c0c28cdf59b8d1c5a582f516 +3eee855312777fee6dacbf993f5c058f355dbde6552dc960d336eee445dd +11d53fd21b745d1e5ec317efbbef25e070d0a36797a6081c356ac2328e64 +e5c55fbc81dc75d9c1575548ece74b8307eef485aa8e28859a2e0435c831 +23a600efb323c362fe9f02407a5411c41a69566cd50add324b63ab939980 +b9d7a929ae4887163cfa7acbfc9fabaab8987a1f6906b9881491cd055b94 +485c968479dbb05b34ed0cd6844729a692978c6928c3392e33e8324ded88 +814cacdac8128e1425c0091a13558100d7cdbed5992795d94d39c32f32dc +621ab6f3b75187a66741f61d6a9c91d791b1cfc3d0e94d4a76302e0c3f2e +cbdc51f14f3251aa5c8bb989f0e13ee500b7b7f2f1e52ca970ad8a7b4b99 +57e93126254297380d67179deb8ff1e99d5cdf7a35c5bb9fa7b402e63234 +78640344e1f10c378ad23c5cd1aa18e1e0b308db70d3a624a455f8e291a2 +ee102ad10776208c2d546cb76d89ca8103a8b95f8acc2d2bdc9791324915 +6c9e05429091071f0c5b76d82c8d1c8a69d840fd460922cd2090624bc218 +0c9391005926a25042a55e322060807363462e1cdeab309033124ba3a884 +1db13f39edae04ec52cde9dbde64ddda1ad805141d4230ec76bd81fd98d7 +0d90fa1aaa26ea551bf687ddd6cdcf3de5a446b266c68434f07d9c0b382d +5816c4e22f22cc03ff78064c6dffb12315c6bcbbf5dc510f5aaabf23471a +234efceeb4aa2f9af9ea787c014c5587ef162fc5b35e8f4c23b168c6e247 +41d33dcc11d2a56d3ba9d8eed6e79aebf9f0faf1a3aeb89d792d69041f0b +b8fadfc0aa090effc6ae5e2f13cdbf54b5bed69b039eef2627769613b6f1 +aefe9b66747fe8feaf7455796740f411a770d4a1764f0483719584880f45 +430e38d3af184145892a08b2add234a3f3ee4ccfc9f6995c02392adafccd +722f366d748cfe9373fbf5f878ed47e9d221fd156bb28369df9e7d2b79da +76120d135ebaf36cff93beb7e313c2b2de7477176fc19609a1b906c995cd +defef08899265b6b8aefb44da1aadefd1c523dce5ca1b84c0c652b3009fd +057789892d4d31764f181754b2e0a62c465587585509989a219711a5e4e2 +5b3b340ca8fdd3f04fef204b1b722b2f6c2ccb00c3cf1a94ba9bdfbfeda9 +e2a062c6f1ced3b8aae5dae32ade1fca1001f98d0ad0e8b36625cc3647b5 +5459552906d8a788eb8bc734ccb65fe9582c71df94fd95d22c5323de235c +28220fb9a2ccb37362174d8cd5922c9e5a87b51d0668555100a33e33750e +f1f795cbed962494a994be7ce8cf71fc58ff4204551b1615ed27cf088171 +fd000b72462b67935961e7c6c3a05bfd67b9ba094ea2c16fdf486da912e1 +e97bfd1c17934535e551cede20c001b5d2adb2be4cbad7d6ba0bdeae4b1a +a739f90293e67ecbdeea4d35825e092697fb05b215083e3f3d6be260790e +2a175fd44eb1c4c16759504827a6eb58a838c4d65fec6eef108495577019 +15740cac164111892e8d1cc447cd208e243a89ab847d8ebf4fb98bff49e7 +a3453facf3b0e8cb67590f390173ddba68324531d2e426aed152e12301d7 +538c1f3c0048a9cc00c009a1a9138460082123209c1e007266fbf236eb72 +21f87d4ca38a0b699e84ca230ffb5095f90a6528bf2a9118f95ac9ab8d2d +ed9eed9b8b27be894b717469758c8d94fa89acc64f530f432d0e5f16c922 +36d6a63410f099c9e909450fd731d698ef658d8ffc1de14817b850814f68 +1a4a9be5cc7a71c381974c249f0b209bfdc2e97f9540c96f57bb4d283622 +00969b82011315289e6a025b137030a0af3b4b42b00fed7cec49df43c59e +3b2495a036dd1b17a8e6adae63bfbbd48807c44b5bbf71813355e1b0e58e +22b6fb88005fc55565be49c17244901b73ef02fc4eb7669be5af22d89c0d +dff0fc6821d810d13e5821d48d4a71d7e463d5b60bc279d0dcf5f8da3a95 +905b56d6f2be95e6d4243b1048e3b662e62401ffaa3bc3f5f90b0854b8a3 +8c38039f61fcb359b06bbb7d59e3b39a295dccd6db9a8b83a6f64ef8dc94 +a77123dd164cfd1c46f1ee51aa19c3d6e7db92a298d10159f2b5eff2caf9 +dc93a6d267fb65bd900d6adf0c6be598050b6d3a9b3a322ab3c9e880d774 +1f58016ff97e5f606b5dbd72ba99252c669209bb556dd5be84fdd7c1ce92 +8a3b3d3aab8d37e6b740227563bb4d60f6bb04052356e1a48d2079feca44 +7ea17fd06f208426d045dee660d1d6460455f8d20dbc5ae64550bbdf60d7 +27d96cc9afef842a8c8c78ea2257e6c6d0d207c80cfe399e8874c693274e +d2c2022d303ca50a70624b07434fb85040a76a823f446c7454dab4f9c05f +10274eb5ba164aa3649d1bc90694316ba5cb3e7df4442e777124cff7ebef +53df2320a0c441ab61666493cb43da46d5711c21699de85bc74359444da2 +e3e397d4c16234f81531505b621aa242a6698886f82b447104b1f1062f60 +b5c87cea9151bb3c627bfa4532b06fd147c556ed8d61ae30a8719dfb8705 +f8a6c74368381403640cc57026d3790c49e2bbd1c0e48285ec6ba44de678 +e3a1394d659c412f09644b83ee1a333a1f51ad8deb4e6d77b3b226ac2c4f +fe653411a7976ae7c4a3cb7df309788da6b483f8a7bab4a6990db74362f5 +bc41d545a320389b2599fd726e426ed9fa2916ece67b058f6a269544e517 +128bda38d117f402409d0d8f8c88ed509aa2ba882e0c579b45af4be80770 +22d7269684eaf0f9afc3054316da6611e3fd260d67fb6fe52c9ade5dda24 +a0050a819ed21342aac9d25194778beb3145f56a66980f620998923521ea +3f957b6ed0c5470734af9f416a16427dd03eff9a0e023452097d4ef936d5 +49a90823cef6de340a1ee02a52851b310cbcf41ae274947a62f9d1d8702a +669023e3caf967204a340694b45fecbda4bf9552f6bdc62d43b3b2c3d571 +9983c182453e22ee34241ab908e667115f7988174684cd70084aefc55caa +f5352a88e9dac45d1ea0e032af61fe9a9118a3931b2050fc6db66ab96a39 +74353b597f34dfd9f72150de23285eda5e555a607d198c291965a7233715 +3f4946a57af0b440ff8567b01a6f46c6d32fea5f8bf57d89dccbab7da882 +ee6c9260e89443b1d7db099477492bd0468850df3db668d741123e7ebe3d +c21748ab4c5cbeb5de33b8963aecafe76bba0c4f6ed8e8263a116ed85e58 +fb71ec4ab0071301be7c7d3afd5fa6ad46c0232807bb7fe129e44bfd16e9 +fd0c8bb5e7cdd86a78b5fb0669093c22eda9151d85b6f58a9c8ead3727c0 +09850bd31a8b4a873d0a506240bb2aeccb8dcb6369532f21d9b967aa8443 +fd6d77cb2d65c4678a5fad188db85940f0a187aa1031dcf5b8e0d0cbfb6d +b3b96fedec5b249b7a69de9b42dfa605bd622de7a220cce9b66e9f3394d6 +13487dc5e82c1e619079cd057b1e19ac05ebdfd7c8bf01c6c66fab49e0b6 +613df9e42beae2f7b9407a2bff8896d8035cea0fd5c11bc5889cb3d90876 +61766138d2625f42d0244adca65d1bc73989328c0eea0b97c7c766285ab3 +351ce2b183f774488a8806c33178090a3808f0ce5e339b87cf7add933301 +ca486742831ca751f0626864ce13172829a8419af5c78794a0eaa17b5bcd +fcb684f7d4bb7af15deb432e44dc7dedf56eb8bea08b46f1e8123a49a349 +a7cbccf833a528f5e22d2d463040e09b91e543a2f33077b3e7b9ecc64f14 +306186cdae1fc317a6ced7e9b4d51a10bbbcf2fadff876b4d9082e3f4aef +dfef230e4232572f4fa33a6e065f6895aa2ea96c5659cb579b023179f0fe +de7ba64bbd9362a7b2b8c4eaec254915629e81d01c839096339b99bc9e25 +84536955feaa52fa20666f65bafd9b2e69c3e8c15d24fa407e7d881679b1 +789a0e2a695d13553c92c0214c9b7562cd6a9a3d77c8b0c2196cef76dc51 +d855c1dac37f96eae4cc7bf07e17dc7c08333d7af33c8b2965ea1f23446b +3c96c52b30ea628ad572694d145b58a606f90b278290297aa372cff56b6f +56f4aad6612eb7c7bd07db4f7d1a70d8044d16d0b5c1605ee02a852ffdb4 +450147b3f9b87d72dc431b34fcdc899462dcc1b6bb6ab1758b6a589e91e5 +8f5196251d00133b43749b7a11fb67a22664c5e38e336dbdeb5509c2d9d6 +2642c07275949df0e2db59314ae0fb34641fc171d3fe1289f919136d853c +d9048ee9db50c699c49e27a8df199590bbc65b23b55bb387eed0c73f2db5 +1cb091f8c22af83103f214199e371f7de1df23f757817200be30610004df +81fe8ed6eba79e856fca21a126ca326ad2f313c16e15754663ad6a065e08 +4050ff005fc899d6e233691b918a093b5f1ffda8839ab23ae66b1bb7b953 +0a7f896ec55de6fb9faf1b49656ff2e57488cd7f1c44114c75f9d571461f +767a6040ffa14e9fb43096f164d60ca530d7cca76d526d1999ac1b52a793 +28651112a65db1f2564ecf90ea6bf2c9ecf515640719c3fb5e36cfc58591 +e227793f39b9d3a9025cb10f324a95c29c488724aa74812366ff0b118fc7 +19f9fd0f202a040be47ec99b46b4dfc3d2a17902a5779c8d52b27231a1bb +5cd794c838daddc3e6824ca8297ba669a818c239b389400faf17aa04b802 +f763029edb9784dfdc42f223e6496a938e613463bf9bbbd59d63300a9ad7 +4e71865cac4b4e81a5864388c3886e70799c8989188341f7d17cb514cd99 +3b211883f171ec6402cc361885f4f4b110757bb3e52941a94bfaebb2faa0 +3e32eb72e25e31abdde82c2a9015478afa0f434ae3f8b97a4bef598d6eda +44ffe1915c26ee0e8339d2d45a6a080550f538ded5542c8b96ca2f596979 +8bb6223e460e857516ab5a3323136ee8fc4b0556a7c39d0cf7acb45e48be +4ae9db325e4750b73289e36a61b301795bdb2ca2a8b933be1c09fd0cd2cb +8677df171d36ef1519a2269b21e4103b2ee151c513df3e10b2a216d6fb22 +18bf2005fa7e0f0563ad96661a7f55e1b5b991f8ca285651b2683c6a7c9d +2d1941374989b06f2e9b42a6af60193dc758dd8e9fcfc7c1aa06eab47e81 +bd79660666defac0c6b9e484df9c17a61ce7a61ef73150e8cd406af6da17 +4d9c2392cc420eddda40f975ffbeacad8ce1b4e14bee29ba8552ff03376f +c034784b38dc1d0ab7bc53943d2545b03d39797af8d58d6dffce56a353d9 +bebc833f04db321ca8642bbb7fcc63ed2349ffa08a33a5d0d78f4fd2c5ea +4258e4671e362036f1f67fcef9d878ae2c203fd9c05200c59cc98633e65a +99d912ec51d6f74500d5358b70e799a6817f59adfc43365d7bba1fd6766c +5c8e76248daf3f01e7a8950fe875d657397797a45e7f99a92887300b6806 +b86db61e03c4c09d6cf507800aeead874a94e6f665746752937214302045 +0b19cfa8db69230517183a03a16e5503882ea1e419c333d3e3b73cef6762 +873ac06bec34c3f736494483442619f5bbadd86f128a5a40b854051893ea +8d31dd6656777ad4ac2572d17c6fb21385b053495d1270e65d78334a4115 +2787ea89b86f97e72718905a11e9c5664837701a3c1c65ccaf26aebe8dab +c1207d5da2079c37883d9235708f370203b3b2a8ec3a5bb35fab93dae115 +aef626dc44b67ca56fac18caf1c22e6fbab93564829a75776630b9c42513 +721ca0fbb0b402f4d1db8f701d2b29fa60162feaa8a167eb3113c6f57036 +e8361357913eb24dd38dc6d3bf4c3176a07ffc75cecf8e5940a310f79a8e +f590844383d631796ade04a91144d073a9413cff34fb454f1fd75cfbe5e6 +525c3bd36ddab80138f6c19aad7417d47df1f1e0fc958fb190a8205b5321 +7c43a4dcb0599be404473d6faebe7240dc402a0e0caa21b56a601b154524 +f44988e5074c71ae8e1948bb2a2ce72fc24cf3b1813cf7408a6b097aff22 +f9d285134d09b7053464259531eb7b270cd5f39f81bbf41a36420f61e5f6 +b429036bbf20e27af1a437becd74c5bbc25ee2519402454fc94d430636e1 +736fe65a643d9b9d21c9a54eac5a8fed51ff60a47b85a0e9423e330e00cf +220c23e056d20aec2fca3e6bc7a61a8366eb940c9bc99fb90e8704e27655 +20335a983eccc7e20b13745c4b4f30a842f1ba64745718c152697c688c73 +6cffcf5cc8eb5756201560413117a45ad3d264291cd51404f98448d31474 +d47d17d201def12867ba679f0e2605de8f3e8135ed0234890cffa68848f0 +6de427741b34c2ea654251ae8450a152538eb806ace3ecfe86d8c4a137ec +c98c6d6cbdc191a5f8f5b5972c70b4896960037b6d4c7c63586a52d5eb59 +47af8c192eb980d0801fa670bb1d08740819f9da1dd9e153010bf9580a1d +0925d8327ea1b88db8d934f40266ddf93e5ea137f267847d826cd7999b33 +c795d0ac05abe2ec1770dd98eea67912f1939118defc9b379e237d6477bc +91ad08e0046b0836fafa1272b0213dce990c90815f5b30d0eb103ac9539c +2f7bd2280264cd95b4be84cbc5139a7628ed211905dcb92cbc3180ac9e6b +b9ecc3cb08608b2395827d5729781dea49d328ba0c1b4cf2cec9f6bbc822 +1f2bbbb9d88f9e7682b9ecc06b9705faa8a90a51678183db1e24cc2c4307 +e16b3c20f08f179ec69df7a8c4261427f5886f9179c493bf2d0ef36640d7 +79925585724aba69df6d1b4f0bd2a356eedfd74a88bea667b102420c2300 +ec420e99b9ce8be1472b617e1255a7f43a0b52f11657f1a4dbb624a24886 +9604fe2062b98f5787d010723e520a4f42a0c3943e151ee627f3d5db90e0 +7747e1a88a53c4784c8d2b042b9c23c9e436d7d88343171161a364cd8961 +37a19582a00d774ef01c7c3fc9e9c7be5074c858d2bacd707a6a4f322027 +137d6ca0421ed9f9c7e7229e867678e5272cfc7156a419e893404ad7dabf +a5d8b6fd0787cb4fe1a901c34dd931f1b64f0c470ff807005fb66350d0ea +eb84ebef2c2399cd14a4454ea5004bddd99988b39c4134b92121ec77faee +55cc716eecc58b594b39c41dcab308efa4458ed71943ec5805dcd0194ddc +1ba04a5d3d42d07ac62a907ea25cd2a7e77aba470324d41dc1a3fe088388 +787b3312f472cb4f23a414fa5f7c7a0cc5d121d7642b5b3f0cf7ca2173af +3f878f374938251feb3ce5ddd2d7703fc79a130978ac516daf70ae903799 +28bea3a4296f48725d578d2e8fb0f932e398404fa8a242024bc011c0ae81 +7b92bb104712253a5d89c543a744332069e33ca08bd133211d233ef799f2 +fed6a20a9073021e505def8b79e1279dacc062cfd4dddc2e8e0a7fda5dd6 +bb5a745f99cccb7ec1df532308da3da0f236c74639c280ea649b2f7ec27d +24221470b642567f3b2e1cd0b3ffa65c5ac986b557aa9b444bf470380435 +abae9b51c6da7ff753810ca7938d8a1c47d2b41fafd236cb5998f3ef365e +1f700bb257679ba3a82e235a3e97a667a6ad94412839c96dcd49dd86ccbb +6df8ad01756b311e9fd57ccd2eb2f19f035e214804e2b77769319a5389c2 +35f3ca2a73c616c9ef0984abcba167d7d652b330c68f4f6378aba69628b4 +2d59eaa2a7e4c782f6eb96f6758d17d35650b15cb5de9bf973b3b6f67c1d +f3285be8322fc2b44359640a3ba5d6d7b96142583a00a9a0ef84fbf14046 +09ad55b2aefe8c5c8f58ed21623bf765f81dbb6cca6d2a51fb7730a14839 +392cad6b47f5e03448350ab36a37d9ff2b9dab69be5196511072b10cc91f +2e6b5160b2b1bd112e6c02d14063a9bb46977b0d4bc79b921fd942f916c9 +c5708e0d133c8309de2f6ee0b1afc996c889c36de20fbbbfd32878f477cd +7735c7c3fa59e9c46e654ea20b4381d9f6c6431082e6918d532bcd539284 +af0333a783c9e7fd4fa1e4da5ce8fea2ea4037644a24532d65fa5c1ee982 +89e4b9abaf71a35d308a9b8c337f70babc5fc8dbb0327143707ca5b675c5 +2d3cf09f7a4f667fcda03d8c82d157e661517787ce6bfb35ea772de13c66 +2bd24b74ff9ab0fbcf6635d8e06b54b5b3125d17ae13d175cb7922338ec8 +9d1159fea2110995ce48f7d2b094f06d11d59b3a64a44a83d48c78855e47 +21243e82d9858401b094a236fa0a90d61863931c30d13b9bf33a35ac0d11 +a999f2b4dfba6fc187f8c235a5217d777a5a97112e7db6a8a4b06b07d9c9 +f41820e233c8b58b9e47ac56ad1ddcc0b35dd03976bc776c6ac3692ec0ca +f8c75ea7825bc84156468ca7b269d890ec9d4a365b0b31d2f6530185d5e0 +2acc3ce14eea55ebb5667067825a8682e135d23c78863d32065ddcf1a755 +e0de6dea7220d1a28416b96db40b1e9f159aeb070c9a9515f301f162b0cf +e32c6c89287de6e2b40458e3393826189a10af8517ff5a10c41c9d05d999 +aa9305a2ee8e7fe46076bc9c5722ee0a140a144ae383e84a8abe70af5d29 +96a0a896cd499caa0ed7867e7c3aac563763216e7769d12218b584d853ec +01db93ca22d0c8d6b286b20b6b26d6ef19f2cebe7030ecaa68d069fac7a0 +09d61770b5e8f83024a99142f59d88297cb8d093992c3c6c11b043b151e8 +20df640407d8bc829bfc196bf2901e63c6f16102d03ffb7c54a7a560f5f9 +5cf8379f4a2eccdcb604bd553e6157b4381940d1b3c768dbfbf2618812f5 +7fbe744b3d8ad680dd9223d8bf2412ecbb614d05b485e3b4669d22b417f5 +02cce2d705c208b15fa83b5be77ccfc1c840f385a58ae49fbe6ab4e53912 +473630e0cfecefab95ebc632a2b10a2103bfe801ca0302542080cfb4cf4d +4c241b1a6c8d28114516e3f1bf39dc02db73e6d9a797279acfd79b02a71b +ae34860dd0e11b18954129f8dd57c039bb7063a4c92f0f6a1e25f4ae59d6 +6c1cc6b73a79d6a56f7f2a8a64d571caa8a760f4f485d770d000ddf393ba +784bb27b781c47678dd78ae9b5d5e8b57d163c42c7a55e4aae22061686bf +aebcede728ff2f65e75955585208c176d100912836b5200a79062d4f09b1 +ba9465b0e937e289160ec543a4cedbbe0cdb5ecfbb4838138ee9e1ac757d +3c5f04fb6b510b389e2f521759e403bfc8ec6bd79e2d40bdd81901c10dd7 +4620acaac9108940daf03af23f09d3c8b785db562b05e597056406557857 +e96fc8bea53c2c2ccd0ea6572abb0acacfe29e737173d665ab6dc2995f60 +807aaa4073a183aed23c26c67eb137c937999fafc63b66a021125e4ee5c1 +a745ad1fff2bd828dcef392052965ce0e9af7a2c88d730fef69da91083fd +83d9fe9f73d42a8dbdcaba85b0fa93b210dbf49cdcbf5d4b69e07375fab1 +a39038cc51f66f0b10eebe0cc61f697f7025d9755830b2d65f1ad0db91ef +ebbfb578053de329935bb28d6ed6c12f748a2f70458990f04d56c35557e3 +8bc5d2e5de7f52bcf00c3bcce091aaa8852d53ac686f8f407baf3f7c8968 +69f3b62f44a5e2291aff9d30d7b5c663658a41add74562dbb0f1062f564a +9b907846291700151de04c1a55cb945eaa2e7a709218ec56d1becce1c0b7 +dc41d5f016ae8080c3b07311590a0def35337fc3c844c0ccd04926be9fec +509b1255ef12f368d20601b1ac8c68b0a935f987a21de0f8191604e921ea +0c04b00dc188fd73499852dbcccd4119ef799472b353be7f7dcc904ddfdb +920839f3d4a13bb1796f2dc886f31217845f8d7a543aabbc720311fd0e6d +a31ad3daa06d5e7e6270a34304f35ef170a7abe733428e96b0522fddbb5d +eb35aacec147067fe066c9ef145246fa3d444d176c274b91fddb8a7bd7ff +7cc7693c25895bf931eb321dc9d79f662a17691f9bd1662fecbcecf6d1f9 +cd8ddcda56d19811f05fa48bcb492feb355b0ec7c04d6046549c56f7799c +2cd0d9dade8809de7d510702e525ad9cc82c41b4fb36218e3d72e905c507 +159076a9c0e4a008ccca17bd594c69f5eee656426f865fc1988d677b72ce +b710b29a0aa8f8337552ae30e93bf7c6e5d013555872dba4737dc5f08c0f +efd428c66fc8da675373f13f89102688977e18e14dedd7f3b676256b0263 +b66b013617d9a026794b0d6040c23c5506a98530249633a6beec46117c96 +ec036eaf6439e25b8e57754af5ebaaf9b57880ad4fc93f002fb03e9fda21 +df4acb78296b0c49a5a852c134c3b10755177a0dbd6c54ea7a2b9bdac62b +5d7f3da649df856478e4baf97899e0f891a96536c283f5c81200c51c6ab6 +77285450c7f7e96836b6da5660f6cb76782ddfc64b6fc348ebc3ba4a46f7 +19176296d8c5a31132b3fa7d935a5d777c1dc84d669d564cb4fd689a38ce +680d0b3b130caea0be43864826d0d154019fd0d865f1c389cd367cb5248e +24640eb6f66603e50581f6fb5aca6cfec1d6dbf4196da10a5e1ebb14e4ca +0251c4c8412cc1673d6e7a9666b04b090567efa0b830d2362fd384cb0303 +8a40290597bdaffe429bb89fb66b9dfcfa92f39d92a8baba7266d144ac04 +f069093ebb3fcea961ba4497d3628ad207e0c8c4fac0e5f3f2a663a8d05d +b6dc33b890ae13d84dce64b495d24cc749b121659373ca31cee09bff2e9e +e5b62e89d5faa4482a75f341dd172500a54b98fc108a69a3ea94db696513 +d4c7691e0095ed3900cd4489ab008b5460b34ae8dedf3721c60de7086605 +6c391137cf23255c565bf11403bdeecf8bf39ad5e4317a4bb37003b2e7c1 +400c3b8ed7f63719bddf07908dc2decdb0f68e8ef722851c4420303f6de1 +b5efc9b2598732fd1f2cbe45a504bd7fbfdafeade3add7274a1e875aba3c +4e0abfc6444944b79f95b5009560818f7a0599e5bab4405378fadfe084f1 +653e5a0166714047e8bd4e4cb116596d8089bae9147ec1d62cd94491af75 +a1743d58bafa11b63b447c954a8d7fe11d39d969feac8fa93c614f97807d +ac62cb7a84a974a0fa555a2e3f0ef662706efcb828ef72e2ea83b29e212d +f89ffecabcb08dbb7119203c4c5db823bf4e8b698b763fbd4d21e57940d9 +1754959d21f3f649d856ac6615eac692ebcbac555f772eb6ba3cece5ebfb +cfcc2f3d8dcad7edc697df93aef762cd47cc3ba9e2cdd10940be676efe7a +a3749170edb47b7562805e3f8bd978b18057c9110ff8d19b466ea238af32 +993e2d3021745b238021f824d887d2e01a7ff12fc6f084b35292f4864579 +406c0f61d0ac7cdf7e4770b424e2ccc22353e6c82bf8ff172973df267ded +bdaabc2a742beea02e35b9b253f98de9ca131f802deee2905ca1a6dc4608 +19a59b4a4265c723007d0215fc8ac2a91ec5f86cd6aac1e370a297103c3a +3cff58c7ae201cbaaa8a12c93e95e73974f9abcd678451b1db02ebb2e10c +c5abfa573a2ea4219fd1851765649318bb556b728d432ec05a86e9894aad +9cdca63d08642655801bb37f28b6e11b958e8e800c8d521ca4aa045fe9ab +ac02dc015d18b1901d519181ef60227170a07f3328a6d5fe4c5aedb35fc1 +3dbe86564a9b1dd4c7ec648880360cdd1742ed4ac409450f1d9681cb5e46 +5edd1de2a2c7f8ed63436f98e849504ae71bb872683ae107ad5df3ca0b47 +a5b79513e02d7c540257d465ae4521cb3449d79c931e2ce8c5b0a0a4ac88 +cef7b9e5f92bf721ad51682d6b6f6c14747f78eaac1891fe29aed4eaf177 +e3d2fc655ae889c0c30a3575a76c52e95db2f6a4d8ffee9518391954b92d +39dae4e97c4022031f8ab390b66ada6dc9ab2de4d1dddf26ac4032981a69 +08f73d34b4849ae28832cddc0dcd116a47d9262b0f93c24fbfdf8a78e6ae +ae3357f3fb89530854257a9db773a1acf5271fc4ca04a06b46dbe661ca11 +9f45e0080cd129e1a7c23a33f1c48af960761b117d9d91fa5a0ed3e47865 +b774a322f7dddfda2960b91fa7ba20c8f9eb213251299ae328b28ef54b0f +55fd54f8047c555e4045cbd70964e1c953e471408e4f25fe8ca7009bfe44 +0244b1e30dff518ea7ce5078027baba4e07ecf0ebecb497b4bd88f1ff72e +b261f6dffec0ed895e237b5608d31ef479e8c9ae9003039a5fe67252ee39 +774e1501100c0fcf154f5c5c81c70539e03118ab91f4ce247f6132d46346 +bbbb126c09d7459c1977e6e367a0c83d14edf7dea081e5f795a7c831fd1b +325b33674ec9c2b68029a0e600746329ea2e1b9bdd5cb2b140468e53c108 +8e8f2567425443f8146ec37101fa4dfccb0e032fff6cdfd76382463551b1 +ae8ca6cbff0e34a3f75ad400a9573217f8cbb00a6d59ff46e48421e97091 +cb17f53f20ebeb89609ea55ed6ba4101f2f3ceccbc7ade21202439ef91d8 +a9a783c22de7e6601b50c4342e094d0eff223494489fa92150425da1b432 +908423fb3f41e0b115ec1ba592a4f920d15610b9fb33f9912aba67912d05 +1ee00a13282c1909a3a56c4ed06f2f4d1739dc296b7492aad0446f87a416 +c6db4d42b504dec3a6756f3d0845ab2d2e151aa5fde12b31a9c3b5ae1cc9 +d97192bc048f00dead66940004281c4d5a92c20b1f77795cb4f98b8eaa7c +be16f9b9d4a34a1a53e0a0deadb4fb4b20d9e8064d3412ea8d2ebd259b8f +2f04bf4bf11a5ab7883c99943d762549c3d5866bb6ed85a0e862eafbcfc7 +03bf4b77cecc0d65bce4df33e0d65456397f231f8cbf66672457cf539817 +6aa5292fae24695009e55904a04588659a3a23fa11989b925705ab45f954 +6f862b0e176fddf75b70d9ef7389f750becbffae25d58a1252cc04a79e13 +fbb6a666fd87cec5562c3e14fd78ad05be28ff3871d6fceff5aa8965bb65 +67ec76d105a6348e915b27767f5010011e80e0e2f9c34742a4eeba369e66 +8faf086a45ac9bcdd76c758db01a78602412a4244c759ece0b963d9ea58b +0efbf4376bf115288803a54cfcf78584c8af80da2a3324096463e3898285 +57de6c6354444b12a74d5e66053f6907c48522cae9e93bccdb4632131add +52eb374213888125de71994c31dba481b70b2e4c1f10b865d58ef09fc9dd +2ca7f69bd2855895256caa5dd6bf7d4d8b341d677c56ca08fd7ba37485b1 +444af8be0dcdb233a512088936ab4d7fc8c03139df396b7408747b142782 +d9406db0dcd31368d2f23ddef61b0da3c0704e9049ccf7f904548c3ca963 +76eadf1ccf77f94c157f5b84f74b0c43466134876a90c5fdc2c53af70c3f +f5c2d13cb665fed9016454bac1a629361c8ea62f4b2399233e8587db6e75 +a9cde3530f20a68ec155d275a4aa6f63aa5cd115244643b54911c954feca +d57be2a6c40f1bac38e393969617b066f7d94e8b18dd80fccd0168d4a385 +f2f1489d1dd41b68d47e5ec66ec568333d1f584e3dca90f1367a990630d0 +14355be7dc45378aa111c319838edd441f15e125f928e044640f25ffdcc5 +c116c3f6ce0d4d3195187b22200808366eca9b508ec45e664e562186efec +a97b22835d384758849605a01973cd9ffc1657b124950c9d9fa3e18b1a20 +7156c4f96f08b87824373c2865845d17a0dda71b1d69f5331c5676d0648b +ca80a7958a2aa034d7e1e9fafead9248e6e64f9ec327c60ae4f724e1fb95 +8a71e82ac3842768b27b506b5982311557432dc3f270ae6eab23a42fef70 +dd0d407a02cbadeb7b8b74a2523cf46a5f61e52b053c2007f75ae053a96d +e00646662d027d93f950e516cddff40501c76cd0d7cf76c66b7bcd1998d2 +7a19f52635c8e27511324aabbb641dd524d11d48a946937b7fa0d89a5dbc +4b582d921811b3fd84c2a432dacb67d684a77ac08845e078e2417c7d9e08 +bd555c5265024aeb55fef4579b46f8c5e79770432c5349d5a65a47ce9338 +e1b599328bb1dff2a838f732852f3debf4bb9b828f9274d03d7cf813b123 +687c5e78a26310d87870bfcb0a76bf32aa20e46f6b2826912e562f503aed +11e427b7765cd2a68da2ec0609259ff14f57c07963d075e96f8bd2eab9a0 +dc32714dd8905f2627c6d6f33563436bda2d7fa9a976f88947b84c72f454 +bf0b66ca84470375d2ff252b4a2df52ab613d0c8ef0465ff1d809ca82025 +c2122a8f44c56ebfa25690bf6a05675ebb8634ddfd24c3734fe8cb32d6d6 +c69c72a4951cb959175770b4286d383e7a3f158450945c8a2ccf7e54fb19 +aa8d2d98a07f0c55f834f2728d89f82a598269750115a02287c4d415cdaa +14e1d9e7032684002f90603c0108dd26b40fb569bb21cc63d0da7e9e1873 +9df0a9c85bc340d2b0940860d95571dc244628c59bab449f057e409e58ca +cc3369f4baa8e53c6765a55620e78341dae06e5cdf2fa5e5ba58634b29ee +ddfee7f78672e55f18a7debbc30862f278f83f4cc123ab591371f548fbf9 +bd24b3453b9b57051c2e67edff2104f3a05a9f0cb7efd81c1b1b0a2bbe95 +21854902526e5d4fa1b3be270811b972e8726623410cec7911c07f871428 +1caaead97c503714eaadb14ae5923f020093722df1b9d9c055d7d5f95af2 +a9fbc5ab6f6c2bd655f685534d7dc5fbb5ebded6ccdcf369bd83c644dc62 +84c2810495888e9d8f464a42228cdc231d5b561c6b210bc493fc1e7bfd66 +5a6c4055a6a629f571f4f05c15cb2104b4f9d0bd1b1f0ab8252da384eeae +f5fd5c663ad7a2c29f65a48a30ed8de196f9eb8ea314c6e86989298146a5 +589f76f12664c8d008228b33144679d16ff564453b5e4e9f813191b6c99e +2680e20a410949ac30691b1428a255b6185b7e3802e8511192e73c376f3d +eb807ad2727fbb4b27538b3213da0746231b1c1b595a958466155835c537 +e0df4a0ef272d4c3f7f2ef011daed38bc58bb0fd7458e48060db98971bd4 +b24bc7bd0de92573a1c7a80a5fa2b34fbe50271dabeb83aaa4235cb7f63d +6a6b399360df8b1235e4e9ab59698930044a98d5e083b5f5a5772309b390 +9e1ff2a252734b32fee3940f0e1ba61f54dd1d3f6ff0d57c9ae75a302d14 +b9dd9034279aaca80b6bd05c74bf3d968305a5046910871223a3ef8c77d8 +25d7e6d3d2809e76064c473d1cd7c05666040b6eba647e34588f49fd70a0 +3c937933a2272c938d2fd3aa8149f215bb48f3bb45090bcb9a6ace393a44 +f1a9bda2ad09a5f566b2e8887880afa45a603a63ffe7c188e3eae926a903 +4f1803368e773f42c7391dff1b9ce8599161515c549aca46aebae7db23ec +8f09db0e0f590aab75e8eb890df354b37cd886bdc230369783a4f22ab51e +0f623738681b0d3f0099c925b93bbb56411205d63f6c05647b3e460ab354 +1bf98c59f7f6c2ea8f29d8fe08df254d8a16aab686baf6856c4fed3ec96b +0328738183dbc1eebb2a3d301b0390ed8bd128bd8e7801c89941485c3c86 +22b5f223cb07dca74f0e8643240044e8c376abbd8c82ff98c6dba9b6d244 +5b6cf4189d63c6acd6e45f07485a0fa55eff370da7e71c26469740a68627 +a3c297d2bf215121fb67815b7b9403aecca10d21e59fabcbe38f5ca66e7b +551b22e28f2d1fd7303d15a42c45bf54b40ef7fc93060ae5164e54f91c55 +20bd303a98d0667a02a900813b260c0343021ac01872fd62cb6abebc7ad3 +a4456805159839ca4a3e35db586221169ded66f852e8974e3815d4d7659f +6a9bb93585aaf264f06cb6da6a26e51683945224158ea69719b8e4e36eb1 +01333aac974db8f84b051724cf245fe7a4c86582b5dbb9a5d9318180e33b +8d92c22c44b0d18f8ca34dfa4ee9693c1a26fedece01635fc5eac1fefa81 +32458254ad46dfdfd2be12a1e7f32f3728f286f1d5d4394424a073696b65 +e3c459aee9310752231fa703faf35e11796c4eeef698f4109ca8c46ee322 +5dc2e3e04fa787188e583321f8410b68b9624ff60679d3f25c13e5ea7506 +a3ce8d0bebb99d9a959ad92d8cf909988d9250b310629903d6bfcad4581a +504b91b2c91889987f36d6fd0be1d0ee5aac00aa0cb48d78a1f7a64a777f +089573ba79452efcc31c8258fb317369feb0d7ccd48cf13da6d1ccb59a4a +48ea0b398e590c1169113fed81639e13e96aa268d99cfdb7aee977fbe85f +f784853a06642b5521ae0a7f610c9739af31ba7a5157ebbbad999e23794a +d2cf25af987dc85dfa29639957cf28e7f2b7671188045130a6e2785f8d8e +30e91f0f68c1cc9f2de902952730003e816e4f5703db7a97b4c566f80547 +42fa77be563ef681a4513b9a68b2b0956551c74545cc9883428dfa72fd5c +4eee93256b26bc86ea34f7427cb0c0cc22c0cc343f739c6c0c46d0923675 +5e04d70587426ef875f8c89ff8492ea23e4e4d763b84a6437a440e69eb70 +65ab6d8cf5f8444a844e6ef3d158b451d121daea2d0e2b423eea24254226 +7eff1b4224c4e80af2a7becac1649e4bbef09f39415e9b1e3750d7ac47a1 +068a4f5ce30840b00574eb4e683e3ec25f6e690feeb0d354568efbc354ba +813ca1400734a67693af127b0f636d58b83e91548f98e3d87da7fd7cdebf +f3ecb4b9272d1c83d4980170378d32f1d98b87c440881af9ec052510982a +0c02ba6743bdc7691a44bae5e044c25304c1a2525cf2c0694494a2e9aa34 +f36af43ab288807ffa4bd418ad51d98c75f2b2f01abfd834d3305682b6b8 +62ef69d05962aac485bb4f560583a5dbb74e967eaf6d299160753ec32249 +bb1d9851d5441cb0c624208e69dc876cd8841a66976b5d7f9c99be68363b +8112d33d971f2c4f2a1feca88ba1a794ddb725c5e2e2c248082231059aef +729bb5fee5006ab8809f63e162fc0743c047c7984a9e6333b433fa143d73 +72d4a74fe37314508e04f54dc7a1445e2d6178ec9c041d0cd4fda5cae830 +4b16feb21f3222261c293a8b058dc708405c1a97ff34eee4ca69ff4e1ee2 +a03380d52297574e3aa50c8afb826fc94a14e8caa9ba89d6e92913be9e07 +bf7ae011e6bd142d8952d9c2304735e875d1ddcf82fa9fc0c6449df2acf0 +d5f6cff6d21ef6b2d29022ed79c4226c97f163284f2311cf34d5b0524a1a +a446645b9d05554f8b49075075f0734b3d1ea31410759c174fcc7305d2c1 +d7128781043cba326251a3375784a506cf32d6a11a4876f85ffa2606fbdf +27dd16d64b2108d808e33c409dd33f6e0c6079e47e7196016f261e824fba +b0e4f91a189747053e648ad2d942ece8f582f052668b63a23a2fae4c75a5 +180db7811aac654270ec6e341126e3561429f1d41fe7ba3f1de9f8bbb8d9 +fc5cebdef869376a2e42dcaa578c0807835e58d75c39f91a83d5c1eb86a1 +b0f7aab991f65eef030f212d38d10b1913bff71717c06c78d9a1be136f21 +4be157ba11ba309326c55c23ae8512646751fb82ae200c06bd2e644bed38 +c7cee826cb587ee8ff378b7fdc00ec316bd4a9c24e2c250cb3d64f8ecbb8 +7f4d81626d7f1e4491908bf17c48c84bb1736693eb4d0fe634484cdd590f +a40ae94d44f348ba683a43004b487f047745fcdfdee2e913328a11a99530 +9bd117e0e5be4fb25d176d59dc2b1842418141190ed9ae1f33e5354cacfd +a5e4bc186119e1461bcd98517e675276ddf0296d3b3cef617dfa36b4759c +944fd721e1bf63d45cea90b5817a40d153a2f779e03487cad3c1375425ac +8cbabf7f754d16cabe45c65f1be4441908e0969d5a5111c931e724537dea +7cd3fbfec9b2f7d3efa747bf586e9218c3106c49276b89fa28f770fa0644 +fe1f3fe3adf07f59c755a5b39a2ac1d6f23c256a293bf3b31b6b9cf4c622 +b188d6e7401c038657c78bfde9ba09f508f1bbe3ed79793772cfc928c4da +519f7dbf3ff7074284437d2de8d7b7c78829642d924abacf353119e9088d +14739935a23667c432806085c3af71ffb7c5fe6b4412b9b1044c1e62ee0a +a5ce7e0322bc65a8c7d874270d84136526e52d0c7f9f93199c6bb7301216 +a19bebcef3c5633f21d012b448d367157ad928e21f8e471e46982bc46a7f +df1bf816a86dc62657c4ebf286134b327ce363ab6a66634eaa2a42e99034 +069fe1302febf06959eab8e7304da4d94a83ac1650a02c38c1c4b7e65c43 +e3a6fb0213e57ac49e58721a4f36996069caedefeb48f1a59303459d5873 +f3bedcdb9d00c1cf31130c27b60928f210e1aa5e1c8e04b86d2049f31265 +9198fa646c53afa9058eb8ceb41bda65f415c79ac92af5790b176de1d300 +f1c06b782d584f458dbd07d32c427d894f84215a8e7819e295ee98d976d5 +644f11920ff2f49cb1075c3bb42b9fe4b561362902f11a75669b7e7c4475 +b65f1ae48834cd67816eb63b58cda2f50bc22eeb0cc965569b476bedded1 +2701668f609393659b266bb0e37bb27afc90bca271366e34754383363592 +0f9a3b508aabfe8deef585b07a992460c592a150b325b1e50e4214a2f483 +e9dfc826c54b488493a96eaa37276f5a9666f0a5388fe388263d2c0cf614 +c6cd01571da4389f01fcdbd0ade1c435d64c5921b5bf7dbebd5268100a03 +1e1abb8cbd83873089a9e08cf80276c7e30d2bb40280278c29fa818eb079 +87623b1cfe13e0b01e27be0a8320b69b5afee820f4705202158b7f3059b3 +655bc28a754d088fde23d43d6a9389da8bc1cf3e8ea1a6f4328c196e655e +42184444d8c0614c7167c91a492c24c8357794c61f5e47cdaf4b38004a5c +8fceaa8151e929328bce1b8f67b22034f3f75e4d105283337c3d460e7d99 +89920c43f5e1449c74ad6ab5ea029cc6e497ea60068451c4ef2132fb87ae +049077a156c868b768df4a4c475a532e2a22d999931c64f8bcc18f51d25f +0f94fbd3e9e6c094f78da062f80c4aa2b86fa572cc469e629deb4ba0c553 +55e8422b562ed2f694d0e8e5540144e30841d7593b255edd4a61dd345d5a +00e411d2c50d64782a3ebedf945fc31c00d2fe4ca800f5aeeaf12ab399db +956362e979bd7ef0787188e43835e5389ac444d13204af6bf1875622f175 +09f32015c28729cfa3b3cca90308eefaf260e3fd9df10f3e76786b8bc0eb +a30e8cd33689aabc55e3ce387cdb89a30573495852a48009cb58a0fd34bd +da911159ccacc94698ffb94c5f45f15ecc9e82365174cefbe746f95eee44 +7a33b4d823487e203478eeb2d8c4bc7b743427778249c56e48fe17d0a501 +7b693509ddfe1f42bdef97aedcc26ceffa9357dd985cdf2c70bbfc987354 +6f0aa7df227ec42f9ca2482f58809e3f9650444568c54d3520bd0a7301ef +48bfebef1fc4332b5ca851fd786c1ece136fe9e575b69393b5aec2611903 +fae6e7a5046e2ff350becb8700f209b1131044afd32fed1bc1297b6a2f29 +6ec3b87f170e92aabacc8867360e4dbce9ea29f0c1df981f6cecc8986767 +0ccfb4c9faeaad7ca9029b8ff0129fec4a040f80ead041b3bc8af7526675 +ed9e13204e64d76440a097d77c535d34165bfe9ffcade530abcc75ae224e +890d5c110004e218bd827a02ac7340e18bf3684c43e664e0a37d5fd4fd1c +4d4489d25a99d542c16e06685652cfa3567da4eb0cb517be1482939da0cd +d0ea3519ad1e51bd9dc7b9077375a8cd3b5de9888697e853bacddbbdd1a3 +0e442e1d6f2d652046821813d0cc0e8f16c97cdd32daf239f5b2b65ef620 +46f6e9821b2e2ec539302747795fa746318514d38bdf0d0e490c00e114d5 +03e7fc9a8fb83b14337a5bb4d640b52630f5450bb3bfcf7cecfbb1ef5192 +ae401265450db197bcfa07315ff95a809bc5fb4249e3a728a817f2580ae3 +50d8d6577f79c883ab4a3119d9ab98219aed0d1e826023a66da814396058 +d95e52d9af8bdbcb0454721f27855b686d13bdb473f650c9865f3e04f08d +b10f5256a3e59bcf16b12a84bb7ef3b370647cdad5929b722a05f5b3669e +14c232bb82fcb9c1dd8155ff4515f4e83c895cafb86754e896f38e5f3beb +5d29f1bd99cb8a09c5e50f412f6d8a773b79021ab2c4831aa663c5defc4d +553616874dd5bd8b75c7a2af7d029aab5a72528fbc4b5ee3d30d523412c9 +60b432434017c4cd68b2062d28f307fc287e11663511d1a6b52143afac0d +ce0f7ba3f326fb707fb8d2c985dd60090e6664f2344e098a7a1a6448026a +2ee651e8141cd7786b6543f512e4c31d25dcaf6652b1eb52706300b771cc +0c49295067befc044ea46341927123ad4b7d094784bda7fa7b568853d0b6 +1e4cc39e1abcc9479f91a2501009ae34ef7d5ff56205cf5288503591cc55 +c48abcc78daa4804549562afc713a4c11152e6e4331619b2e474a25ffb62 +7c46112fa4259f07871f8d6882e9a7ec62d20a86a0c502815d0a8f3f5ce7 +cb4a6a74b6db8e17d54bc919b82c7c729cc05b98855b9d8a0fabd8a9bdfd +4333f395607631f57c0473be0fb290c4f40a7aa6ac49208570ffa1d0f849 +d4871ebcf9ef6f5106301cf54ff8cc9918d6de74d519fccba58bb1c21543 +f3bca9f43c211b2e5c233ff6dff2c9b56d3f656f6070d13dfd0be04653e4 +98c670770e01c07b731ca0e2eb56e608828fedaf1a31087f2d43cb4c0074 +e576769b0830577c86ad5de48ee216df02d7c4e4ec231afd8e76c608fc9d +06cc86f38cf4d839e0a0829902f56cf2f86f08b975a6bdd0642d6b4c78e2 +57cf9a4f52646a952f6a220c36c91db7f44c7f44bddf33328ea8cc01827b +5f2d79e3ee6c514a4f8597a847ef5f32c6400736e6ade28faa7bc6e9c6ba +e4bbff236fa6dd2b0ed23fc77f92649feba149f82488260b0bea2a4fe1f4 +65d96d8c51719e5e10d4c17d1b67e700aac36b1ed55c93b4b2604e72f51e +b30fbf5b64c6fcaaef764639ebd789f82ed354712c7f9fcd1df257e14c0e +8fd59a0eddab684bb1b4176d79b22ad2605bf534e4b8fac2272fbdeaf210 +0424a2c5cc65f8dd5faa13313dd926128ed466046ee94bd3eb41f3ea5505 +5a70603a2ae1981bfae8e77d850fc5a5bf1bacb3df9b7cbce68ce7979fad +a73c2900526b68236c6d37197b0c521c5b1cf5cbbc89238586eceb99818e +aa47ca94ff615233575fe83d0d50d734351e0363030a12300f7b20450946 +17bb209c346ac1d35402b617d6260fce04ce8b3231ab5c05af30b0f3ccb3 +3616d3df334c8d963279537563222dfbb705c3e14616ad01927f952e6364 +4c4b7fa44ac97616c1521facd066aa33b2296dc03682eb6a3b9dd8e5bf62 +53f10667ecb07bbd50553f1b211067f5cf098b64b84d94ba9ad8b146dc9e +8e9be06bc14cfe0945e22fd819856d6996e857c0bb5f292defeb493589f4 +515700753885d61eee1b8c19e6e94fe2302c07933f949d6bf119d207fb04 +dae7bcff7578bf33d77e29611c7cf03b2df12c242827ec4c4e5b5343ca3e +4f7f38ed337583e30dedd78a082f41d60cbad55d59dbba11af1bd296ed6f +e31d2e10d3a8b5ea698e656ff97755a47ddd862d23309e2e6ed3e3e111c0 +2c3a713d782fe301dbaff0a4225f932576622d1cbae40d20f46958298d01 +783851c894f2712bfc4736d3802e548a704878e2d139348671fb96d0ddbb +f56d9349172caef0dfed4b84d867116d91063dcdf9ec401dfe8abb269ee6 +0d646bd12e0752313e2ddc272d9f4aeb9d940987596ab623f9198765cec4 +62f7b6c540c9a70c9a872bd28ea62e056560b61ec51fc68eafe008f20760 +246e06374ae5a6bd2577217700507978811ec29985ab644e474e41e8a105 +295fa67ae05e0739e8c7fbc51104522934942f53e1e1df1ec2a66f0a74b5 +9885cf2c2fad1cab3e2b609f126ac8b7350d5408a7df9ed5c27a10ef6505 +6f0d877cd7bb902977ba93e6e8520d2d018560ec8143876ad0dcb95b173d +af72c0d413bbb5541f14faa57eedb3ac2430e36911d2f486d9ebf9cb6745 +2ccc763e1e46e7a4b8373e06082176a6c66d045e18f90b4b2ad15802f6ef +cf2130cdc627601ecc19887784b6de7fb6a193bc3d057ace29f74199acae +69526ba6f7a2c669593f9d0849f12e37201c32c88384e4548a6718cbb2ab +714ccc917d93b865ac7d7d4dbd13979843f4f5c1f8b937ef12fcdc9aff50 +f09d2625f4367ee70a98772a273d8919952102aa03297e3cbcd876da5abd +2ceb162b8fe1d9a22ff694495528c09a8819fbfb6946ab205d4b2424f6d5 +6fa1c704065cb64fb2aa0fdf291fd5e7daa38667e6d8e889be7f4c453da0 +59c492cd25fcf4a03a6995897145273a66cd6ba999138bc8e2aa7d080f9d +231497ed28a9a27b6b0d4785bfaee46fee71b26d6839f2549a14e7ab7347 +0b6cf368d2d49e74c78d93477828e4582589cb447d795181d3f13dd8ad52 +3c750df8f19b3260c17a6598b406472a7204dd26c5988911ce9884de9a1d +ce33d834becb1dc80efb07f32d3ed6c2a484c5d53746071576c3f67f25ff +1558986fe2dc2265b4fff79c07e3f4c6c0ce8319e04c14728ed722cf214f +65066148bc817753dfdcc0950bf80dc515002e1a92e7d8936e9b3aa9635a +a6d512c68aebc79a62a6bd17a411bba7684e1f06be9bc3d1aca25d50c8bd +1d75597194cf87c9ffe04ff28bea91b5b9521fd356ed9e036466137586ee +f0a8795486438d0d9707cb2854f12963929edac394c562235ca71376d938 +e4e1518668180b857d75318bc22e9f0683749047e7649f9e20b35204b6ee +60c0d47bebf53179a083f0b4cad5b3327a3faf2cf03753e3e46c05773629 +7e9bb305f603369cbb568350b2b5c6d23a35c551e0ab28b082e321ef4ed0 +e2704d35c75b4750af782160c2f2e9aab0e14e541e95b64ebedd66db2c12 +a8935a60177cab634e20a8871a3a72f4b21c3a34d9dac37176a321c2ce3e +e828d140c8445117e7fe4738000c30ffae8e2a48bd618cc8813e38fa0f86 +92ca634d1e56010987483aa0f08980d91528df3d370ac724acb238e141ab +595dcb3da7a769de170edd5763078d1084e2ebefadf8a50a816b50722617 +c9539dbd68d9062b015639708dd900aecf4f15adb36339c05a9aec7403ed +771f9f28c60e52bda3ba6902e06334036c1dfd66d35ed00e3fc0bebf55da +416093b5cf512217c47f905ccc91fad879d63dd1380519a02025ddf15d70 +eaa1bd8cb6be67608fbc5c94796bd09ba35933f64c5e72a26db1ae40ef49 +af5e972fa44660588292b67ac670bf046cb1f5a7a0d73ffd6df862744786 +4a56393b0f1b4cfcfa362c74634713093161b29c94a2526b7138aa92fdde +b37a8c1f30a6b3837d9500b340515f0412e681f5bf36e7869fa157df18e5 +c79df3e6aca924d7b7dd2e0d5b87682d7ea6913b26397ac180fb75fabc1b +8e156ed542b9d8c83079bccd141c187f90d72694de4f6d08520d11cd454b +bd3c2e6d259694fda0c8decc724bdd650163b7f6ce1181590c06de4c0dd8 +536aba318cabf54782c919e07c2ffa1034143175d05deddfcd7dce6c86a9 +ec9bf6a4437da474aac2dbce2c91aedc20043f179d5c9120f3dfb1cf6906 +c27f2ec68cd75035c283e1672ea90d953a23a1515c420b81c3270fa06573 +4d003eca1bb71a2dacdab67e44f47c266c2ea1776648b62bc110671e6eca +4546d3c72c8acd956e10452c32532ed51bf3d0518467fa829efd9c896e8e +1e5c7ff6da0b51e872e403470affc95f25e1d2b9b59ddb0472705e14fdc8 +fc2af16527188508be10d098372cd7eb7d62a85c8d8dd1d0f55ae3ccd0a6 +5dd6bf776dc187bf4de409d5db3fcc5a6d852848a251f4fb4e01dac5e9b9 +587fa8c46ce03689709008b34dfb3dc105def80a1b515abcbe06e73fdf7e +7136e40cc922fe9a9da1726747e84427f288d934747b6c587490734906b8 +a91144ac82a57957cffab561714e1ff5148a39499dfc8cc96bf5d87ced17 +825e8f80cd943d9a73945fb8bc51cf1f9cb39c605491c1bb8f1c4139974a +59471ead310d041b1ca1ecd5e9f92007cd8243cb3fb1ec5256444699a9fc +ed6cb31eaf0912c16fa480a1cb4a8f4a9cb6a4d9a9903d1e2f674286032b +489b8a23ac4719fe435a9fa2d79abdbaba740e69d5ed611421b1aefcd06a +362ddbb7b79aac41e3e90657afc0b87a6e8c57ceef70a628efe19f568634 +50f47b5c6d95870039caa3d07a54e58df064bb5f59dbe9b9a2c7c84d7e0f +32386309560a0efa2cbfa27f861b208b2df4a062ffe2c59c057296aaf5c2 +0f48ffc9ff0692f8cfbd6fc6ed1f3a14537ba40d7267e6b5f69c997a949b +26577a9a99db3f53167355c4967dabd522292ddaca3c537bcf303ce76add +eb99f6664227a94d6a698dd5a5d40008349376067d057e28e55972264502 +e035b1f5e33d7b3aeae016f9be50f2aa09aa138d15d7af3c1ccb805f2d5b +cd4e9b2b5c288b2af4a25abf0a9093749377c9e8232ba1af17962f85064a +23b0a13f11acbb471cc700f9f1b588f72cb63d3d1a95a93502ef74ed212a +c452f1a84619bbdf61a1dc79c0d9ba29c7f19b400f682cf66f7705849314 +f5c8bbf973f2c53bdb060932156bf2c9cd8d36cf6271075500b0e3e6ad49 +958af46a9dc950f4c29f1ab5dc0a85924f7ffef259f778459c80118b1eb1 +ed29208d1145b21b19d62f755de4972c57a09b3decb0a8096ab025fe6b9d +be49ae35394f0ea40d3693980f97f712b27f0e28d8a549acbf1da63518d0 +374941effacf63ac3de0523cfac0dcaeb690de5836741fe58917c7ecffc1 +95e7b560a3e763aa70fc883751bd60ea0a0f893d8e9fe75a66c67e202c24 +84f66708ae74413c0101fe0b5003be20881345d917203b582a247e6c74a8 +1d0479f317aba7b9dbbc0a92e91c51fbe8775a44c57699acc9da84ad60fb +9629929d1edabbd70b4ef9887ce4ec2469f154fada42de54240cf3302364 +7c492ba17e6936a4d85e0751df0945463368a803fb40d8ded22abe118250 +86cfff1878abe5b100bc08b991cda6fdfd579332360f0c3374842edce6ed +e43649d6702f34668a29bf387e647f96d78f33395e8d4b3521cb4fb0956d +12c924c16eee798cde68e319a358cc3524c753177d976d4e14a2e0cb72a4 +80cd87bfb842060b1266568af298bbec58a717c577be73ad808e004348f1 +6aead32a3d57457376ab57197534d6e469ed24474a83618f3ce21df515a1 +22918f4b62c642de0c8a62315ebe02bcfc529c5b8f7c127085c2d819e29a +f44be20fa077ee01a8d427bbe3d97a9d2bafd77f17835279bf135900aee5 +9bc49582b18d468bf93e47ce0bdd627775264ebe9e4172839a444f928580 +8c95895b7e23592b2dcd41ee82e966c26aa2143e3057161511796e980998 +1f2e4ef5868b3bf4576e3546e6407e35cdf14654bcefa7557d09407545a2 +38173080b4771ea52054736677a8d9749a2b22b46b24fbff93c55aa2274b +8c7ddbd751bcaf1df00ccbe1f24a80622aff192fd6db2238db941ec44ae0 +dd73f6b2f80d89bd0aa30c038583deba14913d38a7b61b54522755e251b2 +aeca62033a39ec1143b2b960f9cb87f748428bec3243b8164f07d5ff72eb +f2ef69347bb933241c2401a96ba5ffa3f9ad060c41f4e6bf7280af65293a +bbae49d723dbc4be61d7e13f7a5931a697e7f2c6582dff416341ccf5a24e +9a53686a1e13bbe0bb480c19a4e72a5e477bd29f39dce1a17f63f1e8c696 +d5f8855cefdbf7ce681c7d6ac46798ca9bbdc01f9ad78ce26011ee4b0a55 +786bb41995e509058610650d4858836fcedfe72b42e1d8ba4d607e7ddbbe +3b0222919c85de3cd428fed182f37f0d38e254378c56358e258f8e336126 +9b1f1acd7f387686e8022326a6bbc1511ed3684e2d2fc9b4e53e83e127e7 +84da13550e593bbad1c87493f27b60240852e7fa24392fbf3f478f411047 +3f00a8fdb6dcb8aae629dc7f055d85341d119f7f6951ae612ffa7df82111 +d1ca48306a57a922cf4c3106f0b5e87efba6815f6de4294c7a0394087067 +677889d22a3fd86b0796200300d2716445078027fe0c0b05c86ac80d2095 +ae874324ee6ea3553bcb92fc1522a6d1524f6fa22b71598fbce784a10b5b +61e50307ef4409ffb7b38f27800f2185140ed08fc4ab396050b068025a9d +e4bddcad201e72ed9b41c4ffd4cee743c9c2345b95c5071442defc8ba5fa +9c63c56e209df41d10d93135a8080f7cccacf67e0b0ddb3e0a31df32b83f +290b3c536e9949973cdc80aa5c8a4feee20290a95f68e59f54050192de42 +f27464ee374e4d2451ee8708933b970402c90ca3070843a449d7c3146347 +1efa666a60fd5cbf55a47e4a3c5c318fc1af944d58d32690a2c7eeef09b2 +d94721896e1e3e76e44a8efd524ed5d6f5eb9da093d277441546c6828745 +ad71b6c13f653dd631bc6fc55d0eb4648b7bd9c0eddb13222542f2b6e8d8 +b80bfab4365f4199a41ac690979285d917de79359a183e6fc254b63e6408 +6d33e3c029f472f40742a99f92999f302f79994ffd615f1a848194cb56c7 +12146850f5e400303bf5bcd4e5fdccd1fe2edf5352d525cb15d8327f45a2 +6e3ac276dc8780c65724d28dc6bf9c7c985840070c35e32859168890d599 +a884dc2a90194cc2e9cc6a20c6c0ee11b20adf3aff01db48eb8dba7b0c81 +7fc10cf5a66e8171a2823a4cd22f0e80c82011ae56dd895ae2d3ebe84ff3 +d521c31453e0909cb9b1cf0b030eb6b7059ec38038cae12d0e1cc4b5b3bf +e6c821faac9b8792441e2612aa1ee9318b71f9966d7d3a64abe349be68b1 +744de7b212f6be73a0e1eb2fa30850acc3d9562f989cb2d4fbfbcd5d3ef7 +ba55717da1cabf197b06ee4d8650e968518b6103fbe68fcd5aab70bdd21d +66f09f96208db67c1b345672486657295a39a7fd689b2c9216c6b46a29dd +1283bdba295dfa839a45b86c14f553ff903a6f7a962f035ce90c241f7cde +13bab01d8b94d89abdf5288288a5b32879f0532148c188d42233613b7a1a +7f68e98e63b44af842b924167da2ab0cab8c470a1696a92a19e190a8e84b +1d307b824506e72e68377107166c9c6b6dc0eed258e71e2c6c7d3e63d921 +39690865d3f347c95070cd9691a025825421be84bd571802c85e2c83ba53 +841223435a9ced5dead103b470a4c6ae9efcc8b53331c61d0e1e6d3246cd +aa1b0da347685121196a07e97d21b10ad34e7031d95c1bafa37b4141bf33 +a6be401129dcd64086885f4b5f1b25bce75a4cc8be60af35479509e64044 +d49c8a0c286e4158a5f346ef5fe93a6d4b0a9372233c7434a7a6f9e7ea21 +30c0b4b9f62e3a74cc5d2916ebdaa51a1ef81fceb6cf221e70002a8a3106 +bfbccc2d1809dde18e9607fcaac008fabb72e8c50244507f4013c5a268a3 +6135ead9cc25362c37aa9511589f18d812e6039490f9c599f44e88754ac1 +4f6c1841d570efde27958c7f1b2c68772584e1d12fea252e3a6ec3b051a7 +6faebbf6f5101978e24a9ca927c02065e8e49150a55c64dd30757e8a33d5 +2a788437a9181efb47414dbc22fdeda203d4122137bd045611f68314e12d +1d6a5ec270c8919562c03e3af7b0e0deceeddbdaf3eab8fb5632e44dc1e8 +d46e2396b0236a46659164e33709415e7b347f7f7b87a9224a189ddf5178 +2cf66c9d385470a51efc88696176f6d3ac3b7b95fa074c981194e22981f5 +1d925f980393b7102f1f836b12855149ef1a20d2949371ddba037b53a389 +7617c257bbdfcd74bc51c2b40f8addfe1b5f8bc45aa4d953c0d1d5f4091c +6af796af6513c820499969593bfd22f8c6dcde1d2ee2c0ceebb5bd6a1ce4 +5fa61094e932b380cee381f4485e39b4b1797f2a7d8d90bcbf89b9cb1006 +2d50fff083743bf318157caac1c0179c87c03a2857fc002979e7cc97feda +966b09ceb761d3f55cf07637256c6aa8b8e5cb6aa9739452a330afbe7082 +975ee39fad5e8106e8ee05771157e92d99003533d922ccc37add065b6236 +7613d039741f99edc77c230fe8d1baba720a185186662376b947bbe1a686 +4b42c61ebe1abd40d890751ab8945c629de3b6d2a49809dc693f9e397097 +cf1e568c258081242460af2de0ca44b7ba2734573967b3bdec0e5e64598c +cbf41e630d821491504f414d9b54a3100dd5105a141cf61bd3ec41b67368 +c8cd366c543754ee800ffee3d19c9cd0d408cc772da10e4d8134964b0a61 +232e2dfbeacd0fdee12792504bb327a2e1fc44127f8577ca51d380a760b3 +740e6be46455cbf3917b90f0dfeadaa25d5d9f66cda43ebf9f75e0191a06 +25ba29666bbe8678822a453d4e876bad4a6b0d4b6cf98feb60339c9eba2a +dce4ef7faba428422c503d0210dcf8d884ca9f5094aab9f3b1a2238b569f +444748902907cb0d9d7ca33fccdd0cd29bc68e44f7bca5092be6272bc949 +baae5af92c302bb21f91b6ea8463265680f7c16f45d8ff35392a10eab87e +296f3af4478032b5b021db8510deb617941130d45c46fb3647d94b162fe2 +2738766fb6d76a06ab6803818b27c5ff4205ba668f95b5ec5ce4ce6da545 +c13ff56f417a4e0b3b8554a1e2a985a167e168adc8c4db28a601a80ab451 +91bf32acfd8d25c39c2f17fb3bca1296d3d160f25b43b4d6b94f20ffe012 +b779339b12860dfc897b366e3d400e756f4f9f4d2c86fb9d94c11ebd1450 +eaf720056e2c39529331bdcb104d113b42c94af2c6a5035750b7ae7fdcba +b6116d74bc07a11d4357ecf73d99221dad5cba4a7136425c2a3ac0e092fd +606a4ab722195e3b7fdfb5a5e3ccbb85fc701c42bec43b54e964dff3fa04 +193043eead7681cedae9cce6919949ea60ef5630c4b9263c8f98b4bc74a1 +63ccf3d0a0bc1deff39b800ac90bd734dda7ecdc73169ad77e129887db80 +7a253f8807a422eda8a16c9ee9bb8fc0942634bfe035dac9f7e36d09844e +39477c043399db4d07b3617da9d6eee76d0fde9201da98b906050748b68d +8c944ace3c96e90a3c2b63eae27b9152cb7274fa336866d71b65a57f1bc2 +bb1f482a67f3993dcb3ff24abb0223f9a026c81b2b33127a1dad8929dec7 +5d46bdd790eb1addd771c5c3965a2f514d3a128117a44560cc10a729bade +4e6c86de7c09a39602235c803902e34f5c176b18e127d71a011dd9a3a61e +ebfaa4a4e2a5651be6f4067e5e09bb4f3514d67c2129e4d3ea9568661138 +1e45af07bd84f883c70577a986416747f3bd8d1bf86d3d7b07e8a350899d +3c2dae237bd5ece45faba7a0ba30fcda7b7eec9fbeaa5a94620686d1e403 +1cd2512e8d89451c7bd8eb432c8862023d66f3f9fcec0d47598e2df59525 +d673a5ff493d458748cd6341f161a0a3e8996ca5b496508578fe4f653924 +2ae28bf4b7397c02b726fd5f9d8b898938bb668a546be6e42865f4f030d9 +5faa289eb24f7b8e249b224a95a2245605d67417a489626df7417855b8d3 +1c0043cadd2b461d32e1b39ccf409757c37b68f84e752bde6b5bbb847bf1 +57ea3434802def983d6ce5ceb3e9fbc4911b5484e99bb94dc3f383e50672 +0e85a91ed378e352838cf02921ee0ea94be01b5a60f9b1f58fcc1b4f527e +43725de9b9dadc3ef462fa279bd7138095d4cff2a0563039f71e383430dc +f628dc9611b2e3db08fb2da1d5383dc1a3c784e1e64541fde1d9d7f42505 +de96d3d0a401099fc2879af0293b0eeb143b78cc221f670c0479bc150047 +0cacb9a282e334e428b527acdfbfc56e6aec8d4d60745c1dc000011b6248 +d9ab4a17dca7cc74e17d33c0641710b02cb1edb0addc6be214b17e9f845b +2d9c8bf03c19e131e00f91f2a393b5f2ae7c3d4ae9021c4d7891d84d5067 +377ce92836e42eacd7e540824f7ac95360ce116d41d17a50748748971c82 +27f089a22ee0d21940de854f737547b73c7517addd9bdaab425a6c2908f6 +87dd990d6cba4d84308bdd4c4435a6480ecfa1a14daabd4d8e2398178e48 +de28b84f7ce4b61d2e6e64fe043c29a941f6de7621ee6f6d8b506221df05 +db238b8fe4323cb5f259d4d3d9c94d4ae1ca37d6c34345489c0284171346 +e9830e2e3c6c167238a7ffe0989d3eac870cd44102cae139469b9d909b5a +9c34792f693ac94ecd35d2277080e30a2d24b50391b6f2a3d3b6c81f7ed1 +a7b218903e7fed7a63269e27d793a2e0b40320ebf447c71f36d40dee002d +7257f43c8add31edf2c571123e46fdb413e007cc89e99b6f98d77ab38bff +cf140f787e45ffb2c7cc4ddbb59a4e32dfc36e2875f204ac851d757c1236 +12deb31324ea4c201d27fdab46e9f3988ad2bcfb8e9cfa8c487831a9b0c6 +60b20fb66b4c77f52359ac96f3b3d189aa0571c1c53db06ddb10f08882db +0b1e93e9478d4c75626c5fbdbc6044c4d82684b310ab2af144d12bf36f1a +c0bf6249d1da9ab319453594cb19d0e93c4e047fb49229c0cce76d0cece4 +2e76fabd2425382afe707db032cf617b046a59a2fc1bb3838d98fd5c8053 +ecb918bc14762e4ca45027623988f434ff4cb08bc9bff5d7de21940e3e03 +1ee042d9c30662aa76f96213fb5a92047af60f320e4660eadd1ec19d0086 +072f2202af5f219725f81882f10d1e065a8035a9946d0ca0e48a5e7dcf61 +0283b834eda01e7d94b3453830daade2aa6c947989b290c02ade0d7b2620 +813ad177ed82813b6a985d5c0a2d42419bda763d409da085936e33c817ae +68e5467eddc30be172de855a0f7f5c527555b3f4d942401b450f08273b1e +c5b5352fdb8562a71f276284cf7c27537e628f94bcbffe8d669ea2645752 +60830f1e65e83a2204cec393f6d92d4f61f317471b4b93039d298ca2cc94 +eeada0140823a2bcd1573e732e7b4bde7368f2ecca5961ad547f554ae989 +98d87b7e5d07a85c382bcea1693a697224f41eb8b406bc6a0c3eddfe8b5c +f25b11c3e4bd91ea7d6274cd6b3ee7b8f18cc3fd502a324c645568dce9e0 +d43caa61f7306fd5488fcfc439d85f8160ebf0ac90fc541f9c74d35d7833 +09309807a639477bb038200738342e50136dc64baa7cc1b879c61f7e1b90 +e1f2bd4f6e54c4dc97b8e4adeb102979203a31fe26a7f58c609915a95abc +4acc263179423f8ab16b04272d5592fc536f29a45cbcdbe15890f119ca9f +c7a52eef41dfa5c4fed087eef8e698ba738e300bd58f2a1a10da1198c1f9 +b60e2032f8384a86aa84027df21cb87977528e3bb9bea1e3a6879c56402e +a29063afc6ac0194f4944433f9a5872cf0a2a741382d7f3c0ca7817d5d7c +4b8bf53af0f18b1eb54480519cebb61d983157e039b13025e7980eb36f54 +3451bbb84e470ffd0f98eba80c74f238729dd6278294388a2e06de68a719 +47b6d478c85f124d14aaa835620e49b7f5a4f21347302c0f0864f7ebaeec +d0831c36187cbe9c848736764a31056d2cef27c07cca00033dcddca9a2f3 +b9ebf28e67257b69cd38bc23c711b6a2f6e4dda9bf5a19da275e6a8d683c +723bfbb95a90a344a6f421f0b67ae84c74652288b0597e4c86c28f73808a +77455f2948e8df634c2d14f221626b019033f9230c9167982cca9ae6dc37 +aecbcb49fd9fc1dbf2d11bba7187888721bc42a7f47c23e07d2fc5a7a91c +0dfe255a7f9d17e69af1618502a6b90b1dd748c7eaca1e1ebe8b861b04ff +e5f628f47eb4e7e65311037d7a5713d7cc3552dc85f452ba74c4f12aecd0 +d72892c940c3325640d62fe3bbbc71361dce6d54766e1fb99dedcb2d19d2 +fa6fa21f9116e03952ebbef659816a62db51a9b5b3916ff818518774ccd6 +79d44100d7236f211f36fa80a4cbafb3db76ba1e7e7f12082b0140eed2cb +5e793e24501715c6c170ad4f856a4bf16bb10210025156e635264d3cf18b +1fc1e8cd2fcfdc2ab1a24af9087975bfcf6fb703fb36e288e58d0d2ffc98 +bb4318001d931ad6161dcdf8984e6690e0f6bb07af81bf07445f8f57b355 +6b960d24e7cd152708489e4d953ab6a155a757e002ead97585e6c5333d7e +5aaab2731f047f3490432e0ebf3d0d628eefa8c1f665b9c86aabb0706639 +5bc372e16378f0d9b439c98e7bf87be73e934995d58e4e70d3ae9a5b54c8 +87a19f2826a772c39d41805c642354d9bec75b065f148f7c1e435dabbeaf +e4a5744e3f2894a928121ab069bffa3218a106a9dbb83971353a7c7a5616 +d9da66fbb908173f9b07aadcbd4d112cc353e7b70476046ce5a92e86eaff +4eec40acc840005f51f55c9f5874216851e9cf3fa431d95d3032e779e356 +4bdce33966a3a798b170a06c4cc9f73700224c858c36bbf2d0326c337ce9 +46f69c19a84187fa50afc5b36010f9a7612e3a25e846d49bb907af9505e7 +d8c78748d7dcb501bbb3d6603e829deee3784f2f3ca583d3738d6d2ecfb8 +eaa887103606211a3c1b5cd74a3e0e96fb57da91baebaecd3669661e7b1d +579ba41928a40a7028acff6cd409e601d23ff66ff2c8acb12e535360d727 +60d2e988d801930e0e9443d60dcb9f378fa75d58d73e6a3b6e5b26407c82 +67d50ad97787f8a9b91765e41552283cb67e43e59bf71cf08b9755c8ce47 +0cf374832c72d1e9702b55bcfc8b5a4e966d5072fb2a72a2108574c58601 +03082ac8c4bba3e7eeb34d6b13181365a0fbd4e0aa25ffded22008d76f67 +d44c3e29741961dbe7cbaae1622a9d2c8bca23056d2a609581d5b5e3d697 +08d7e369b48b08fa69660e0ce3157c24f8d6e59bf2f564ce495d0fca4741 +c3a58ec9f924986399480ee547ad1853288e994940bd1d0a2d2519797bf2 +8f345e1bb9cbf6997dae764e69c64534e7f9dd98f86b5710ff8b500e1c4d +f509da50c64e213ebdf91978553a5d90908eb554f09b8fc2748c9c405903 +e7bfbf0ea7e84254fb6735f09bf865244238e5fed85336c995bc3a3b9948 +947a6eb95db4cd1b64c0fccf82d247a2202e9e7eef5a550557625a0192bc +8bcc9e461e52833f6b8729ccd957d5c4b6e07016e864fc02b792c7400ace +d0a8f43c755f87bba6e5c6e1022416e5454cb34a19865d951f7aea527760 +53658cbf306ead832244f3062c39a0a121a1157a8e47008163c5bfc88197 +be16e9a1ba26a035a16dd38cc28dffb666dd4ba7356c66b7bced9e26e905 +4ce25f6d36607d8f5dda1e21ac96a815bb2989f01130ba1aca9aade554fe +effdfef5d6b0d2a01aad92f599f6a12e121010ae6acc6f150f19e7305271 +97da761b07530ca19b84b119e5edca1fad18462143b8913d6b3f6864b713 +7a93bb9e1bc29c09d660704e8d8292c61072ebfe35c354a2342b2458a353 +31d043874380d439388e46688a53bcfe01bc190ef1a6b5dec9d40aafe822 +261b28bf3e2d76f3dc4302506ce3387b4aa2a51cd4ba1faa2ed1fd7df664 +6772fe9f83d253451eeb0448b444b8ca80cc7cb653c2d1eaa0de6f2b1c72 +47e6d24ae72e620e200aff83a557a1aa7a0ce0a9cfbbeae03c31d8cbf1d8 +20b53b688ed2ffbd83418d743ee31e3d62216ac7be6c12bc1917548cf670 +d69fd2e78d9f7786ada0ea30a6f6d9fbd1f1406337151ffa1d3d40afbe03 +728fd1aa2fa8a4f075796b9de9586b71218b4356fb52daa01d3c18cb75ae +d4d33fc809dcb6e3dcf7aee408a0cef21353d76ed480bf522fdfe86e0e0a +b7d097defcb793057f0ce98ea4989a9b6787b14029a4bf10315a2557149a +fe9c91e7d825f7518b343fb556f0177a8f6ca08fbda9913d52997511590e +b9942c9813b4cf4d4aae4919401f2fc11fef0620eb5c40532cdb22d5fad6 +919a3a710de6c40d54993b5386636499c866938e33bc703a99c73adc228d +95cac73ff4f4a275c04d0d787b62c6a184dacc4024d23f593e7721be232e +9882fb738160e52ab905f0ce2c76ae6ff2c8bbe118a1acdb3b464178cf01 +94bc6a50df1090e9221be11e49f254b06c3236a31569b947ad041d1c6b55 +bfdec3c18c791ace0fe2a59504eef64a4eec4b5c8dd38b092745e0d5ad29 +276bf02c419c546627672a5764a4904635bff86fd0781d36fbdf13485229 +71f355de2b0ad250052f50ad70f61afc870ac7a816561d3232b73360d4ab +2727b2fd045f254c782bb3f1f49d94c6d625047071b7e32da5c6d21a86de +9283fd632074430772bfbd85e0c9ccab1dec16bbc049c3e223bec1b65c8a +9e98cf58b30a74f74f1a842dc91e30c023498e280ac55edd58f4cc731d81 +e443d9b9efdf5fea63c9f357320e01b8740eedaeef2495cd02eb2f338b3e +674fb074cc497d7b1937b188da857c2c230e9a931cbc00c85a7a36fa80b4 +56588e1bbabbe4ef429a6aef9bd4eb89c5752421bd049aa13f4dcf9b51ce +2503e90bc118fac78a25d187353d6f5d496cd6130b337666f49619cea985 +dfbeb7e49c67c1e0f0f8e9ec8ba14624ed0982dcbb69415e4b3c8ddba140 +397eb1fc1ddd36c94c374f018873ba41109e45afa51f0e691157d5958c06 +26fbc0903ae25e47ee372389cf65472a3e4d9769550bdc42c0b72f9a297c +d5d3c16ec67e06036e740ab664abc9f10b9499269b73ad3678daf4474329 +c2c7252c1f0df1e3b5e8f198dfef8325cb1e7e8057897a3d7fb5bb5858e0 +cfc0c115bbd7362d8e8ee41862af6eeda681cabbb06f72ebd2ae0b0be45b +a9e1be83f1da30687a655e5d148fcc17d9f53b760810a565f6d2f4cd5da3 +5434116edef756adb4d3df544a1de593be988f2bb8d36c34deaac7d9dc15 +cba49764f1e03aa09fe21fcd7c74e3d6487ebe219569e019f10dd163046b +c1a3cb2bcbaa8558197cb2c18709a998b4efa8ab8c9a71d2ccf942c17662 +1b88dee6b424165d6ce10ac48375e760983818e0085276b1674dd41042e1 +a01a8de111c903f74834199b3230bd475d92c6226ef74eb1daaec3475a6a +fcb47644a17c7e390ee3b16bef1c1ca6c55eddc44fbefbdde525921b3047 +0d76817bd8ac724739a8e743eb09cf78e88adad527d4f115b8a32ed4898f +45bab3eb802b8168aec061e3ecdb026c056fb9efe7e2df48bd516ccb12ce +00de08ed8be4ee0c41f40f4c8f64483e0ade90a78d6d4fe9203fe0b97c60 +3b2f8882bc15a212453c691c52d00fae8a3a26934ff8acf68d4352eef75a +0b10d938e55b7333dda2db0296a69e9775bf82b1aa6d684fd9080fc1c11f +ab4369c7a95a9504063db900a6e345bf6dd99be041230b2e60cc86b8c345 +1d84a9c2cb4ab6d74d63dd43dc26eb6b384f5222796d4083dcc3e1651548 +d9469f09a33b213a33ac52a6a2e23802d8f8a75c01a607940daab0051410 +73a88130bc192f303616adb113c0051b65e12086cb319c0a5323fa7def40 +402f5f87a3b2c2cf0e92789985f6775ac2743e1ffe2d0668291059740d45 +43bae7a2897e5e658592bf5a72966097742e0702deecb0cb12499eab701d +34ba37a08346217a415e44297a181bbf3744f0a49230ad6f030e11462be9 +afc2ae14e0587bc02311b48b8e2122c28cdf14414f3680fa52dbbb63b17f +6ebe4a1204f3c5d6150cbf89a8023890383153838d4dde77d4c8b1b78823 +8918c564d3babfe58eeb154307dd1997f5ab7105426e35c279008b2677e4 +695c60f956b348799c04b734338018fc27f7de7ad9d73468fdbc5283bd14 +c066ddad9a3562f16baae15d72d7bfcb409e1c874e9db1a8cde233b282b9 +6e76e9c08d85ddfbd3cce7e64104d0b0e95291bd91f405ff82f41601ee20 +8471e613fbbee67f269e4e954c36d1d18ca9880b7cc2b08fc990978efdc5 +1d157deefedaa765c1e26ee125d4a2514a41a3b95e9151a824532d7d6486 +35ad622718fe71219a697e94c2e64f26424cbb767acdef5cda70e179cd29 +b7e318d1c6d3ad26fd5fdcbf2fc221301cc1f10f5ed86b40a1a6bcc01c90 +eafd65183e75609610637b99fea57885efe76437df02a2ffc21223d039b5 +74955d9a54ff41980eddaa8768c5ad883a0c9150877392b990d63c6805db +7b8d6ab1358cbedaedb6feadb0ee4fb8f9c1ca03a3e755a74227a8930bb7 +2ea0a00b48fc626fa14d7d48624aedc31c556f44e982f3ccbde7ee735f73 +629ab1b65bcbcf0a3586a920477e8c960219802fcb1bc3a179032b324f8d +c424899b38275886cb5bc771f26a0880767d49cc23426a40a4b6ff8fe48f +d747565fc537565f6d7fd08706accc60f5fbcb45bc785f45ee9b0812366f +ae71b23ec43f3549c8224d78baf18719f05108d5741e681457ead8abc050 +462481771a8dc6cfeb98956e163981a98c59ab44d90e9c3a946c453b5071 +db0c769f7fb5144c7ab0c9ef1a6db1addcde1d4ae1daee1b4035af256a04 +df53926c7a2dcdb94caaf12f986e20929ba4e396f3aa7c93a7abaef1294f +5f13a0dd3c3aaa8fb38da3e15daa32163b7437af683b4f5e64cb14aebbde +8c69ed2e8cdbfb213fc8129af29ca2c06c8f85a5038d688d1fa5d1b54ebe +4dea81a49ce24131f8e6702e7aa4e2cba078d5dd373f894ccb275f49c690 +1dc772e1d2f5fb3fe15dbfffac62c87110162074eb72ae4e5e446bf7e650 +a554178d0d64d3c07f330f0d99e99f2239cb1597f2e5f443854cdb0f5fab +b28fe62f22e7f3419d017980f325351bb04f8f3c3dc57fee03cc029bd29b +202308d5a800ed2d500d41ace8e54e2557bf25b627883beb8118d800eb94 +f4253f855168f7fc8a2d29c5fcb76bb90a6c4e345722b8991a854047f46e +4e97336be85470b6be2b9ba573dbc4967ddcdbfc3b6fc35b0c7f3f2f570c +55dc3fee6d80bc6f46cc7e4d86a0b86f6fa61d062e213d9e442db63fbf11 +d03165b44572096995ed342893bb672f6bb55ff8fed944667995f0f89a48 +a904c47420f32afd14129c6e2bedffce1f07ea69d550b6909bb5beb4aa08 +b0b44f35e018ba5206fdb4df0228462c1fdbb95a429e53eb27bb1b0490db +f07202c3608d0f4ce08570e3d6aa3d4581c569b57bd8c1ea0e4ed3fc5497 +e316ecec06e6be582d9170d426f6d22d8c7287b8219945c124941ca8812b +e97efd9105eb6999edc0665016633b3b48820df736125b7c76c9f3a67d93 +8a2a0a6b743fd42aebc46a0249be459f16811ac9eba7b63bad7c2e88f175 +0eff8da5faaab5659824f9d19b3225aad2ac17c52c523414d3031d08a926 +30abf474fe02a32b44d3b7d9fe0c19aec16ca6d018b71d9d395ffaea0788 +0d4501d7cdf0f7077a2d63303d09083080d67f1f714a1b271dab9fc9866e +4b0571a171eec8a4e351ba2d02438cd108a33b1106acaad0ccdb051061ea +7f40543748115f29debfb4be4b42cae8762d62114ec6f8ef68c478a8e05d +ecfa18b0368428efec9eafb2353f95e3d71e1636b9d9f94a77e692843255 +698576dce13b2b858d2d15ee47cdba3ed08d64b77ab46dd29bba6aac2106 +ab847de378cccdaf35c64e50840248915f4fc110992c493cb1b9cd0b483f +0f1abf5e9b018210b477fea28234ffbe5e0bbe01338e0842a89f1e00a0ca +7cdde0b2d7c324d5e17d8d3415ccad703507497ac95360ce660b656e5f66 +72a2f50761f3d02ccdc1d5692d7797699b8e2147cfd4817c81a432ff6a5f +39cc54927fa146cbed56a55f85f123c0a94b7553a8819b329d9dd122c502 +94e3f6314d5117db89ae7597c4691b6c542979a1ca3d26a8e23d3eb698c7 +1841651e08ec771cfb974d6613f2143872c739b62796bd0a45172530793c +28d93a65b59f79c245248d2c09428657a35b0c0e367bf7a4a4f0425b3f4b +485d9f402e164328a4b963f456829a39035c00283d2e4fcb71a42da6d42a +d46cb751287de34e6519c60bb3f1a6ba91f7bfa21dca96ee712af5681701 +18ece8a0535d9ba1dd4bd835e004a2f38c5ba43c9b30d17045e5649fbbac +188922e442182d4bdafaefb39e00106a5a7765f3d67850471e3629e526af +8691f935b57bd38465665204a214fef1006ea37dc0781073ced5fc042781 +93650393c3cadfddedcc5550ed483bb6355f54600e9758e647f9c9711f1b +e7df05d0e50a698615307c18f6d4886f50188011ba499d03831185915f3f +77c4b9ce708d78423b110776aaaf90396be0381616d1e9b0c1dcf68b6396 +82399da2a7323bf42ae5347599ef4ae9e5c135522c5ecb87e201853eb899 +db60d24acad17d6b7c2c7ea4dc221f3cb6d6caacd1ac0822ea3242ad9b4d +d15116c3874e3012fad26074a23b3cc7e25d67ef349811dbc6b87b53377f +0cf972040a037ecb91e3406a9bac68c9cab9be9a6bb28e93e3275b177cd5 +0b66935cbe8dd3d6a8365625db936b2cfc87d4d6e7322df3dbe6ccda2421 +a5e5372566f626a5e9d8bc66959e443286f8eb4bcfdeb6c49a799f1efa69 +63260d0ea2d51260baba9207fb246da927fc4c89e9c4dd5848fd4ef6f81a +cd836f5f06ff0fe135cafd7ab512af55a57727dd05a5fe1f7c3c7bbe8ea7 +e6680fcb3bbbee1cf2e2c0bba20185f00e2dc3afd42f22de472cdb3eaa5a +ddf8c6fb3682eea5548c51ddca25ca615221127b4438ea535ab3089c9ed9 +b971f35245cf831d9461a5da9d57bc4e5606d26535a7414cef6aee2a7b95 +bf2276044818ee0f3b0a16532934b8b745d8137b42ec2b28fae7d55fc02c +9ccfa4e0055f8a4be96e1e235c01b8b6ad509b832a3e90161e0a449934e7 +4be973c939b31cbc19dad4c58e9be89d242f0ce200548cdd4fa2081ab3f8 +e01f358d5db24b7a50eb2096d833378921f561f132cd7988708ee10cffb6 +2256201801c667e176b1dfaecde9756d725bef093457805e16f550e8a7de +87ecd46e5b09646b73ee74f890a36867636911e4cda2c46a40e7d57cf297 +9696046614c85b1a47ba55c60544ebd3ad7d750d003bda56dd7eed8c4702 +f8b319aaeef9d3cdc59b3e63ee93c6e1e857af273eb90909ecf36ef4c276 +895c78aa762e5376c5c542f854fba864ebce56e4b0207091139f053c2c08 +3b7ddcd0a9909b52100002bc3f8c47bcb19e7a9cb58b1ac03fee95e81195 +072d3aa7c8079632725f63425a3550a947834d29ac9a26d0774e90248e18 +996731fd9aa53ab62b40ce557d98e874b763d9d629a173f0c7babfc00ae7 +82daef5f00cf3608ebeef403dbbc19e16a1d160b889f4a10359d9eacc19d +7b5f126b31720dce7fc35ec861dfa56ea23fa18423ff4e8fe6e53fc6ba16 +b95a2b5dec00f614e4f835281ee0b4bf549e7e882689e0b445dd46fc40c9 +090e5575fa2c34b02a51ad0bccf6a7bb83ca3b929285e5e9fd054b72c47b +733a66c5abda526b18b2e49d0746e067e63b948a45eab2f4221c5b62ae21 +a5d9d7cd8aa9eeb49588891d22c56b14b55ceb6488f02b73ab3b7f6c5555 +b75452594658255e4cd58ac4815f2e1bc3888c6777f62aac2f0a57d416c3 +765c991f0f9a33d888aeb2d527b482c042ee23783a04a73ad13dfc590a52 +f3116f8296cacc7ab29b7d87e7864561a5d0a12bde2d36ee697064f41d1b +ca6ef2f801caab5295d19bf4c02b10c19f73b44635ba48a0806b967d7dfc +ce9a4850171a78532cb30020c0d66b3b1e7c75eaa7894904c181a022e8bc +9b2b8ef1202f3c7d36bcab4742d4a4761bb55b64da0d99685d319f5da8fa +132be6c0483f50e2657ae8af1e28f969440d6ed43eb00e95fd9e1cd490a4 +8646f6d008598751f7a41b43fbec7770fe591012b6b0c4ae18775ccc7db5 +de0ded2dd53e82c89648d46f0d0cc5d3ac5aa104239608d512a4353b9547 +04fe6eb7e73d718323cf9d748b8ec5da01ec9358267de12cc22b05ef0312 +e4b6ac5dbb6d06d7f2d911f20d527f504d62547aef136834b3695df8044c +383b6145e824d3931a602f081d9d656f84987a1ef121772f1f5b37a116bb +d2e77d4ccda01411545d24e15ce595db4cd62ee876b8754df0b85b44e011 +b82d76ce45795e6c2c58be8690b734a8880a074f303a70da4a1b086a6de6 +56c02cc7a4c25258eff18cb0fd868214bb46f972e26509f868d065b3cb14 +1c316898cf22293391bd7051ac3a6927aada952a8fd0658ce63357c07f34 +acbf8c99a5537da0023e901f0eb5547e1b466b7d982c8c539798b76ee2a2 +252437a81a37c3b63f625172d682eeed0b795860b2755f020ef52a138353 +003c61be2052cdd7d73b2cdcd26b127660a7b22fc51a6a2f6034f37e3e46 +c1d7f83f8b28c7c965993abba1d358362833580d9c63fa85d4cb949f97de +579fb6807b95a58b78f596db50055947dd0d0e597d9687083e9bc0266e86 +90b884b27f4094d8fb82ffdbaac4d580340a9ef8aa242be87e54b601af19 +87a48d267c04e371ae77163ebd0de3f5297b1060442ecdeac38334844e38 +0f294d4be73935fd8a38de7fba6d082c3d9156d7e88f2cfff0459377cbb6 +041f37a7e05010753b98e0b67d5827aa312129bb3c3bd883c12323756406 +d555720da8a0bb30edcfa760c01ecc2ba3b15fecccf5a10e9f358822e0ff +b64178fce2ea6a1105bfb72df0e4bc499b207ae26b8ea960de48e7ee7010 +b4e671dff795e4cdc5b43e81b1604d224f0616ae311f1208859c502c1a10 +940e7b9cd11be728bd3a0c8005ae23aea32c1b642812198a6f1aed32cb75 +97152b1340dd35ada1b81051e393d38f3740fa9523df6a83b8ca7dbceb33 +6e299b54cd998d4dfef804733c76156585e42b7284cbcc4047ba6b290efc +aa60953e98cd2b4bc2893857fa6a339f820142a52ccab0df09a2709df550 +f22e5921cbca408e7998cc1cccb8adf6d8f8b71e6685ae59d290fa33f5cd +664d73e434237424060f634262f04e9a71a977556e93b692ddc3aad26d92 +97dde71e4def64932151ad572af6e681082e9944ddbec6e7a8bdfd534233 +9ca3106ca1ccc80eab14f1655978b137fad8f399df7cbfa2d7d3d9675e0e +9afec37369a8ede2c93145ab3f42a375926946680c215fa16bf7416fc892 +bacd806cd424b9f85b47802c4336918f7486af2a03bf0d39b10169d35494 +419cb1ab7b8f407897f70c18303e91563b497d70b7181ede6aa0c3efe089 +ca6135b34dd1019b298e3677f8da61f864a67023c31eaa716c40cf3d397f +9a1209564c9ec759c37028079661d2a56374203c78b023ec61340bce5d96 +e477a4f77e5c0db7c0d1257b4bbbc6f889b17e6eaab045b8adef6f931e4d +0795583d60a6b7002cf61639c6f930671f3b8ac05a1c4e002f4bfc50d8b2 +3029fc4dce1b602cc3a5533336271bccc226559ffb127e3a562f92f89824 +552b9a70466d5a3c74ae515a222b109d490f26e8fc2d9d72bc8af6d1dcc7 +80463c7af81993bac2ce4aece9d95ab736b1dc73e32d1237bc8ec2b52513 +36dbabb4ecc7ceb5d18b02043281eb9a3bfdf19bc4853c9b1722ef1cdcf4 +fcec534923db2e2653dc48545a9850c0ac2e4594abc9f7d18a0bcf2fadfb +bf085d465a4d10528312f5d790eb9511ca01061c0d94136b99a043bcf278 +c18223b1e0f1cc062b32b79e28dec2dc59a0aaa4b5f3506923c83e6a87fa +08a1d941bb644c994491cf7f3b0e2ccf6c8a8ba89376f76dfdb592374f93 +528e78e31e0b18719346b9f1486f652638e3120687774030444674cb0778 +96385c41f6566819652d825dd58f9a4308ff79b45d7828dcbfebc406e40a +c46e866cb0e3e97d6ce7fcac19a9d0fe39bbde66c5f0cf775eb3b1e6d7e1 +1f67e7edb3d5c4facc85c916bf13322b56a0414ca27d145cb740fa2c37cd +8c142d9301f1ac3704cf6a8e93973a07fde5a331cf0cbb370c7ba555de61 +18a6cea0ecb2c0e37152390cc57e2e4fb3791ddbc383ee26b6f4006d0d68 +4880888011020f856a9de47f45440f127cf27ccaea7d40a3869d39ec7dec +ebc06382d294717644b6118354e15544fd4c6d88df9245c9a83b30e6ce09 +e2498dd1df488a019b179cb859889e6ad2838f749e3b038b280ebc8d5c3a +b03e8f15751214691edf0f86281e612d7ec0773c8a5d2b433266402df62f +fcc06879ca196aaf1fc73a5f01ac46b44d6cbe7743ae9a862c20445ae2be +1544f413d010280cc2941900bf3c42ec088cb21b44a915bb810e7666b545 +5324465c5943eedcef0c09128a995f431382e2062f5e39f4338c8eba1bca +e553cb60bb8f3e5038ac8073398c49f06dc734b18afa7921ea0d455e6e73 +db8ad9f77fb5ba6c28af6b4f18cbe46cf842c82d6c960be1520a5fd929df +ac7e00ede976fb2be0a07f659079a421fca693de89ce9b8fcb42b0176d9d +f3ddd58f921e13e216933d27b49d175b423751c451be7618eaab054d3b8c +23e8dd6fd60182d61e9b5c86b3b764a29a62f913ee7524d8cb33737d7224 +d95dc4bb8c2ad6397604a0ffecc8865adcb540e5da1cd769077838515118 +ebc9f0b988545c1881dd2e7a8fd73e11bd7ae9085fb4d45526b23a346b0f +e4281ee3d588106db5f7c386c488d8f2f4dd02d4c08e74c1034f987a44e5 +d39fd07538de57a42987ce290fb2f6557e8b5cbcaec168f5780927226415 +1e11e3667d33b36a793aa53e9e2d1102c9eb30cb3ba0ebac953e0227fe4a +3d3c0eb57e4390c3d35db0c41946e45be2830a1ae33fa25cf2c7c9cb4550 +ce9ff6c6e3d628fc7284daa6241604c90dde6339b7f7e7df3733416cdac8 +e5291357e4983d74d3582a490438a7fdb0af97001a31990b1de68e6adb48 +917daa387e647f9f13312db57310c7dedc2a2ea80800b4f4bbaa99c6b7b2 +7ac8345cb659489307e2565ebfd17774642c9ae5d3c18068dc35170c7d58 +4cf4173f1baf98137fa249c81f3347e1dadd6b1ba0f50c3b64c1eab183a0 +937b0f7278eff101e5267fa6480da7d602844416490c2c2c7eb0d44ac8f4 +75cfd611db5ec268db07c0b3608825c3e12834a2b2efaf5e2723c5199c42 +6011cf22e64e4c0d31d563f321097935ea0c6fcbf5acd3748d90079f6ab8 +687288dc55df29fe7958f566b27b73e2ea30747247f7a2b2add0602c7d64 +d23f52e7c96748e6a54ee8c4629b2aab8882169653f0ba7f05236bf14364 +244720f3259cbed73a318b29e4a9305deb65a2c9dec8a9d0f9a9f6fae541 +83e0f4b9a9a567057a1794945168dc23cec25d1c02ea9242c9fb6d8fc11e +e8874bd80a5226373ae87cea91853d0625c777ceb1f5a6f3debcf2f75a61 +460c7b4067f568ecd01f62901ade8bf8fbc5db9c6720420496f0cb48a002 +99870773c2e7b12e83987a5d0290d9bbf589ac889bf7d4334a5147187a7f +71008f216ce917ca4cfba5347078f354897fd87ac48af6a6c62711d2eb3a +5882bf3b32c0f1bfda976f850c9dcb97170e78c229a27fd5e292d161ece9 +a8c47a223cbdc28e24f79f6429c72b5752a08f917feda941582c36d9acb5 +748c86072858d053170fdbf708971a0bd5a8d8034ec769cb72ea88eb5cd7 +49f35be6ee5e9b5df6021926cae9dac3f5ec2b33680b12e95fd4ecbf28eb +a0503c10c6f2be6c7c47e9d66a0fae6038441c50e6447892f4aaf0a25ccd +952c2e8b201bb479099f16fc4903993ac18d4667c84c124685ae7648a826 +6bc1701cc600964fdcc01258a72104a0e5e9996b34c2691a66fa20f48d7c +2522333dfdabf3785f37dd9b021e8ee29fa10f76f43d5f935996cbf9d98d +92d0a84ce65613f7c4a5052f4c408bf10679fc28a4a9ff848d9e0c4976bb +dfdfb78bb934cd72434db596cb49e199f386a0bda69449ce2e11e3a4f53d +be134c6d7fe452a0927cf6a9a15b2406f8bd354adcde0ce136378baa565f +b9c51a03b1fbe1e166a1f92af26bd9f072250aaa6596a236ba2d5a200c90 +a760ca050421abc78223b2e8b2eea958ab23084fa1947574e846e48aeb12 +26cebb8b5a92089e9ea771557599e2fff44d75bcf600e76ae7289ba98cf3 +98208c5104562834f568ebd62801b988b0a9fdf132b6564566103b3d2d8e +6a099b7fbad8a13b8cd7f6729bb6651fc1019e66c4bd6ff27410bd5cdae7 +4010bd68b066bffdb4fd5e3dd9cf7e1a1353f7a4c5157e3ad508f4ca0259 +9761b7cdd6a81b3560b8765be3b0432fe4c25dcb4001b00c7fa62874f681 +ed22127dc3974605a05be8d8fcf9701f859ffce4dc598091891ab7596ac3 +4cd851ecfd2dbbaa2f99dac376f7bb40703fd0700d7499a7c24726bdc9bb +3b88c6a82e52686c1ee945d8825092bc81848a08722ac5a1d24353f95ec8 +18f3fa487d9600318091b0ae9874b42bb3cb683a2518b18cc1bd86c6e5e8 +3d37c14ef4fe0c77b03a3314995b1e7c1066b98c4375bd1fc5fadee1b024 +7ece4f95a0f59978d543910deb2e5761632c74c508269c4e4b9e315bda02 +975dc771fc30c8164b9df9172a4e571d8ca578cd2aaeaa0dd083e74cdc2e +d938b984b96d76a64b8c5fd12e63220bbac41e5bcd5ccb6b84bdbf6a02d5 +934ac50c654c0853209a6758bcdf560e53566d78987484bb6672ebe93f22 +dcba14e3acc132a2d9ae837adde04d8b16 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +cleartomark +%%BeginResource: procset Altsys_header 4 0 +userdict begin /AltsysDict 245 dict def end +AltsysDict begin +/bdf{bind def}bind def +/xdf{exch def}bdf +/defed{where{pop true}{false}ifelse}bdf +/ndf{1 index where{pop pop pop}{dup xcheck{bind}if def}ifelse}bdf +/d{setdash}bdf +/h{closepath}bdf +/H{}bdf +/J{setlinecap}bdf +/j{setlinejoin}bdf +/M{setmiterlimit}bdf +/n{newpath}bdf +/N{newpath}bdf +/q{gsave}bdf +/Q{grestore}bdf +/w{setlinewidth}bdf +/sepdef{ + dup where not + { +AltsysSepDict + } + if + 3 1 roll exch put +}bdf +/st{settransfer}bdf +/colorimage defed /_rci xdf +/_NXLevel2 defed { + _NXLevel2 not { +/colorimage where { +userdict eq { +/_rci false def +} if +} if + } if +} if +/md defed{ + md type /dicttype eq { +/colorimage where { +md eq { +/_rci false def +}if +}if +/settransfer where { +md eq { +/st systemdict /settransfer get def +}if +}if + }if +}if +/setstrokeadjust defed +{ + true setstrokeadjust + /C{curveto}bdf + /L{lineto}bdf + /m{moveto}bdf +} +{ + /dr{transform .25 sub round .25 add +exch .25 sub round .25 add exch itransform}bdf + /C{dr curveto}bdf + /L{dr lineto}bdf + /m{dr moveto}bdf + /setstrokeadjust{pop}bdf +}ifelse +/rectstroke defed /xt xdf +xt {/yt save def} if +/privrectpath { + 4 -2 roll m + dtransform round exch round exch idtransform + 2 copy 0 lt exch 0 lt xor + {dup 0 exch rlineto exch 0 rlineto neg 0 exch rlineto} + {exch dup 0 rlineto exch 0 exch rlineto neg 0 rlineto} + ifelse + closepath +}bdf +/rectclip{newpath privrectpath clip newpath}def +/rectfill{gsave newpath privrectpath fill grestore}def +/rectstroke{gsave newpath privrectpath stroke grestore}def +xt {yt restore} if +/_fonthacksave false def +/currentpacking defed +{ + /_bfh {/_fonthacksave currentpacking def false setpacking} bdf + /_efh {_fonthacksave setpacking} bdf +} +{ + /_bfh {} bdf + /_efh {} bdf +}ifelse +/packedarray{array astore readonly}ndf +/` +{ + false setoverprint + + + /-save0- save def + 5 index concat + pop + storerect left bottom width height rectclip + pop + + /dict_count countdictstack def + /op_count count 1 sub def + userdict begin + + /showpage {} def + + 0 setgray 0 setlinecap 1 setlinewidth + 0 setlinejoin 10 setmiterlimit [] 0 setdash newpath + +} bdf +/currentpacking defed{true setpacking}if +/min{2 copy gt{exch}if pop}bdf +/max{2 copy lt{exch}if pop}bdf +/xformfont { currentfont exch makefont setfont } bdf +/fhnumcolors 1 + statusdict begin +/processcolors defed +{ +pop processcolors +} +{ +/deviceinfo defed { +deviceinfo /Colors known { +pop deviceinfo /Colors get +} if +} if +} ifelse + end +def +/printerRes + gsave + matrix defaultmatrix setmatrix + 72 72 dtransform + abs exch abs + max + grestore + def +/graycalcs +[ + {Angle Frequency} + {GrayAngle GrayFrequency} + {0 Width Height matrix defaultmatrix idtransform +dup mul exch dup mul add sqrt 72 exch div} + {0 GrayWidth GrayHeight matrix defaultmatrix idtransform +dup mul exch dup mul add sqrt 72 exch div} +] def +/calcgraysteps { + forcemaxsteps + { +maxsteps + } + { +/currenthalftone defed +{currenthalftone /dicttype eq}{false}ifelse +{ +currenthalftone begin +HalftoneType 4 le +{graycalcs HalftoneType 1 sub get exec} +{ +HalftoneType 5 eq +{ +Default begin +{graycalcs HalftoneType 1 sub get exec} +end +} +{0 60} +ifelse +} +ifelse +end +} +{ +currentscreen pop exch +} +ifelse + +printerRes 300 max exch div exch +2 copy +sin mul round dup mul +3 1 roll +cos mul round dup mul +add 1 add +dup maxsteps gt {pop maxsteps} if + } + ifelse +} bdf +/nextrelease defed { + /languagelevel defed not { +/framebuffer defed { +0 40 string framebuffer 9 1 roll 8 {pop} repeat +dup 516 eq exch 520 eq or +{ +/fhnumcolors 3 def +/currentscreen {60 0 {pop pop 1}}bdf +/calcgraysteps {maxsteps} bdf +}if +}if + }if +}if +fhnumcolors 1 ne { + /calcgraysteps {maxsteps} bdf +} if +/currentpagedevice defed { + + + currentpagedevice /PreRenderingEnhance known + { +currentpagedevice /PreRenderingEnhance get +{ +/calcgraysteps +{ +forcemaxsteps +{maxsteps} +{256 maxsteps min} +ifelse +} def +} if + } if +} if +/gradfrequency 144 def +printerRes 1000 lt { + /gradfrequency 72 def +} if +/adjnumsteps { + + dup dtransform abs exch abs max + + printerRes div + + gradfrequency mul + round + 5 max + min +}bdf +/goodsep { + spots exch get 4 get dup sepname eq exch (_vc_Registration) eq or +}bdf +/BeginGradation defed +{/bb{BeginGradation}bdf} +{/bb{}bdf} +ifelse +/EndGradation defed +{/eb{EndGradation}bdf} +{/eb{}bdf} +ifelse +/bottom -0 def +/delta -0 def +/frac -0 def +/height -0 def +/left -0 def +/numsteps1 -0 def +/radius -0 def +/right -0 def +/top -0 def +/width -0 def +/xt -0 def +/yt -0 def +/df currentflat def +/tempstr 1 string def +/clipflatness currentflat def +/inverted? + 0 currenttransfer exec .5 ge def +/tc1 [0 0 0 1] def +/tc2 [0 0 0 1] def +/storerect{/top xdf /right xdf /bottom xdf /left xdf +/width right left sub def /height top bottom sub def}bdf +/concatprocs{ + systemdict /packedarray known + {dup type /packedarraytype eq 2 index type /packedarraytype eq or}{false}ifelse + { +/proc2 exch cvlit def /proc1 exch cvlit def +proc1 aload pop proc2 aload pop +proc1 length proc2 length add packedarray cvx + } + { +/proc2 exch cvlit def /proc1 exch cvlit def +/newproc proc1 length proc2 length add array def +newproc 0 proc1 putinterval newproc proc1 length proc2 putinterval +newproc cvx + }ifelse +}bdf +/i{dup 0 eq + {pop df dup} + {dup} ifelse + /clipflatness xdf setflat +}bdf +version cvr 38.0 le +{/setrgbcolor{ +currenttransfer exec 3 1 roll +currenttransfer exec 3 1 roll +currenttransfer exec 3 1 roll +setrgbcolor}bdf}if +/vms {/vmsv save def} bdf +/vmr {vmsv restore} bdf +/vmrs{vmsv restore /vmsv save def}bdf +/eomode{ + {/filler /eofill load def /clipper /eoclip load def} + {/filler /fill load def /clipper /clip load def} + ifelse +}bdf +/normtaper{}bdf +/logtaper{9 mul 1 add log}bdf +/CD{ + /NF exch def + { +exch dup +/FID ne 1 index/UniqueID ne and +{exch NF 3 1 roll put} +{pop pop} +ifelse + }forall + NF +}bdf +/MN{ + 1 index length + /Len exch def + dup length Len add + string dup + Len + 4 -1 roll + putinterval + dup + 0 + 4 -1 roll + putinterval +}bdf +/RC{4 -1 roll /ourvec xdf 256 string cvs(|______)anchorsearch + {1 index MN cvn/NewN exch def cvn + findfont dup maxlength dict CD dup/FontName NewN put dup + /Encoding ourvec put NewN exch definefont pop}{pop}ifelse}bdf +/RF{ + dup + FontDirectory exch + known + {pop 3 -1 roll pop} + {RC} + ifelse +}bdf +/FF{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known + {exch pop findfont 3 -1 roll pop} + {pop dup findfont dup maxlength dict CD dup dup + /Encoding exch /Encoding get 256 array copy 7 -1 roll + {3 -1 roll dup 4 -2 roll put}forall put definefont} + ifelse}bdf +/RFJ{ + dup + FontDirectory exch + known + {pop 3 -1 roll pop + FontDirectory /Ryumin-Light-83pv-RKSJ-H known + {pop pop /Ryumin-Light-83pv-RKSJ-H dup}if + } + {RC} + ifelse +}bdf +/FFJ{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known + {exch pop findfont 3 -1 roll pop} + {pop +dup FontDirectory exch known not + {FontDirectory /Ryumin-Light-83pv-RKSJ-H known +{pop /Ryumin-Light-83pv-RKSJ-H}if + }if + dup findfont dup maxlength dict CD dup dup + /Encoding exch /Encoding get 256 array copy 7 -1 roll + {3 -1 roll dup 4 -2 roll put}forall put definefont} + ifelse}bdf +/fps{ + currentflat + exch + dup 0 le{pop 1}if + { +dup setflat 3 index stopped +{1.3 mul dup 3 index gt{pop setflat pop pop stop}if} +{exit} +ifelse + }loop + pop setflat pop pop +}bdf +/fp{100 currentflat fps}bdf +/clipper{clip}bdf +/W{/clipper load 100 clipflatness dup setflat fps}bdf +userdict begin /BDFontDict 29 dict def end +BDFontDict begin +/bu{}def +/bn{}def +/setTxMode{av 70 ge{pop}if pop}def +/gm{m}def +/show{pop}def +/gr{pop}def +/fnt{pop pop pop}def +/fs{pop}def +/fz{pop}def +/lin{pop pop}def +/:M {pop pop} def +/sf {pop} def +/S {pop} def +/@b {pop pop pop pop pop pop pop pop} def +/_bdsave /save load def +/_bdrestore /restore load def +/save { dup /fontsave eq {null} {_bdsave} ifelse } def +/restore { dup null eq { pop } { _bdrestore } ifelse } def +/fontsave null def +end +/MacVec 256 array def +MacVec 0 /Helvetica findfont +/Encoding get 0 128 getinterval putinterval +MacVec 127 /DEL put MacVec 16#27 /quotesingle put MacVec 16#60 /grave put +/NUL/SOH/STX/ETX/EOT/ENQ/ACK/BEL/BS/HT/LF/VT/FF/CR/SO/SI +/DLE/DC1/DC2/DC3/DC4/NAK/SYN/ETB/CAN/EM/SUB/ESC/FS/GS/RS/US +MacVec 0 32 getinterval astore pop +/Adieresis/Aring/Ccedilla/Eacute/Ntilde/Odieresis/Udieresis/aacute +/agrave/acircumflex/adieresis/atilde/aring/ccedilla/eacute/egrave +/ecircumflex/edieresis/iacute/igrave/icircumflex/idieresis/ntilde/oacute +/ograve/ocircumflex/odieresis/otilde/uacute/ugrave/ucircumflex/udieresis +/dagger/degree/cent/sterling/section/bullet/paragraph/germandbls +/registered/copyright/trademark/acute/dieresis/notequal/AE/Oslash +/infinity/plusminus/lessequal/greaterequal/yen/mu/partialdiff/summation +/product/pi/integral/ordfeminine/ordmasculine/Omega/ae/oslash +/questiondown/exclamdown/logicalnot/radical/florin/approxequal/Delta/guillemotleft +/guillemotright/ellipsis/nbspace/Agrave/Atilde/Otilde/OE/oe +/endash/emdash/quotedblleft/quotedblright/quoteleft/quoteright/divide/lozenge +/ydieresis/Ydieresis/fraction/currency/guilsinglleft/guilsinglright/fi/fl +/daggerdbl/periodcentered/quotesinglbase/quotedblbase +/perthousand/Acircumflex/Ecircumflex/Aacute +/Edieresis/Egrave/Iacute/Icircumflex/Idieresis/Igrave/Oacute/Ocircumflex +/apple/Ograve/Uacute/Ucircumflex/Ugrave/dotlessi/circumflex/tilde +/macron/breve/dotaccent/ring/cedilla/hungarumlaut/ogonek/caron +MacVec 128 128 getinterval astore pop +end %. AltsysDict +%%EndResource +%%EndProlog +%%BeginSetup +AltsysDict begin +_bfh +%%IncludeResource: font Symbol +_efh +0 dict dup begin +end +/f0 /Symbol FF def +_bfh +%%IncludeResource: font ZapfHumanist601BT-Bold +_efh +0 dict dup begin +end +/f1 /ZapfHumanist601BT-Bold FF def +end %. AltsysDict +%%EndSetup +AltsysDict begin +/onlyk4{false}ndf +/ccmyk{dup 5 -1 roll sub 0 max exch}ndf +/cmyk2gray{ + 4 -1 roll 0.3 mul 4 -1 roll 0.59 mul 4 -1 roll 0.11 mul + add add add 1 min neg 1 add +}bdf +/setcmykcolor{1 exch sub ccmyk ccmyk ccmyk pop setrgbcolor}ndf +/maxcolor { + max max max +} ndf +/maxspot { + pop +} ndf +/setcmykcoloroverprint{4{dup -1 eq{pop 0}if 4 1 roll}repeat setcmykcolor}ndf +/findcmykcustomcolor{5 packedarray}ndf +/setcustomcolor{exch aload pop pop 4{4 index mul 4 1 roll}repeat setcmykcolor pop}ndf +/setseparationgray{setgray}ndf +/setoverprint{pop}ndf +/currentoverprint false ndf +/cmykbufs2gray{ + 0 1 2 index length 1 sub + { +4 index 1 index get 0.3 mul +4 index 2 index get 0.59 mul +4 index 3 index get 0.11 mul +4 index 4 index get +add add add cvi 255 min +255 exch sub +2 index 3 1 roll put + }for + 4 1 roll pop pop pop +}bdf +/colorimage{ + pop pop + [ +5 -1 roll/exec cvx +6 -1 roll/exec cvx +7 -1 roll/exec cvx +8 -1 roll/exec cvx +/cmykbufs2gray cvx + ]cvx + image +} +%. version 47.1 on Linotronic of Postscript defines colorimage incorrectly (rgb model only) +version cvr 47.1 le +statusdict /product get (Lino) anchorsearch{pop pop true}{pop false}ifelse +and{userdict begin bdf end}{ndf}ifelse +fhnumcolors 1 ne {/yt save def} if +/customcolorimage{ + aload pop + (_vc_Registration) eq + { +pop pop pop pop separationimage + } + { +/ik xdf /iy xdf /im xdf /ic xdf +ic im iy ik cmyk2gray /xt xdf +currenttransfer +{dup 1.0 exch sub xt mul add}concatprocs +st +image + } + ifelse +}ndf +fhnumcolors 1 ne {yt restore} if +fhnumcolors 3 ne {/yt save def} if +/customcolorimage{ + aload pop + (_vc_Registration) eq + { +pop pop pop pop separationimage + } + { +/ik xdf /iy xdf /im xdf /ic xdf +1.0 dup ic ik add min sub +1.0 dup im ik add min sub +1.0 dup iy ik add min sub +/ic xdf /iy xdf /im xdf +currentcolortransfer +4 1 roll +{dup 1.0 exch sub ic mul add}concatprocs 4 1 roll +{dup 1.0 exch sub iy mul add}concatprocs 4 1 roll +{dup 1.0 exch sub im mul add}concatprocs 4 1 roll +setcolortransfer +{/dummy xdf dummy}concatprocs{dummy}{dummy}true 3 colorimage + } + ifelse +}ndf +fhnumcolors 3 ne {yt restore} if +fhnumcolors 4 ne {/yt save def} if +/customcolorimage{ + aload pop + (_vc_Registration) eq + { +pop pop pop pop separationimage + } + { +/ik xdf /iy xdf /im xdf /ic xdf +currentcolortransfer +{1.0 exch sub ik mul ik sub 1 add}concatprocs 4 1 roll +{1.0 exch sub iy mul iy sub 1 add}concatprocs 4 1 roll +{1.0 exch sub im mul im sub 1 add}concatprocs 4 1 roll +{1.0 exch sub ic mul ic sub 1 add}concatprocs 4 1 roll +setcolortransfer +{/dummy xdf dummy}concatprocs{dummy}{dummy}{dummy} +true 4 colorimage + } + ifelse +}ndf +fhnumcolors 4 ne {yt restore} if +/separationimage{image}ndf +/newcmykcustomcolor{6 packedarray}ndf +/inkoverprint false ndf +/setinkoverprint{pop}ndf +/setspotcolor { + spots exch get + dup 4 get (_vc_Registration) eq + {pop 1 exch sub setseparationgray} + {0 5 getinterval exch setcustomcolor} + ifelse +}ndf +/currentcolortransfer{currenttransfer dup dup dup}ndf +/setcolortransfer{st pop pop pop}ndf +/fas{}ndf +/sas{}ndf +/fhsetspreadsize{pop}ndf +/filler{fill}bdf +/F{gsave {filler}fp grestore}bdf +/f{closepath F}bdf +/S{gsave {stroke}fp grestore}bdf +/s{closepath S}bdf +/bc4 [0 0 0 0] def +/_lfp4 { + /iosv inkoverprint def + /cosv currentoverprint def + /yt xdf + /xt xdf + /ang xdf + storerect + /taperfcn xdf + /k2 xdf /y2 xdf /m2 xdf /c2 xdf + /k1 xdf /y1 xdf /m1 xdf /c1 xdf + c1 c2 sub abs + m1 m2 sub abs + y1 y2 sub abs + k1 k2 sub abs + maxcolor + calcgraysteps mul abs round + height abs adjnumsteps + dup 2 lt {pop 1} if + 1 sub /numsteps1 xdf + currentflat mark + currentflat clipflatness + /delta top bottom sub numsteps1 1 add div def + /right right left sub def + /botsv top delta sub def + { +{ +W +xt yt translate +ang rotate +xt neg yt neg translate +dup setflat +/bottom botsv def +0 1 numsteps1 +{ +numsteps1 dup 0 eq {pop 0.5 } { div } ifelse +taperfcn /frac xdf +bc4 0 c2 c1 sub frac mul c1 add put +bc4 1 m2 m1 sub frac mul m1 add put +bc4 2 y2 y1 sub frac mul y1 add put +bc4 3 k2 k1 sub frac mul k1 add put +bc4 vc +1 index setflat +{ +mark {newpath left bottom right delta rectfill}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +/bottom bottom delta sub def +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/bcs [0 0] def +/_lfs4 { + /iosv inkoverprint def + /cosv currentoverprint def + /yt xdf + /xt xdf + /ang xdf + storerect + /taperfcn xdf + /tint2 xdf + /tint1 xdf + bcs exch 1 exch put + tint1 tint2 sub abs + bcs 1 get maxspot + calcgraysteps mul abs round + height abs adjnumsteps + dup 2 lt {pop 2} if + 1 sub /numsteps1 xdf + currentflat mark + currentflat clipflatness + /delta top bottom sub numsteps1 1 add div def + /right right left sub def + /botsv top delta sub def + { +{ +W +xt yt translate +ang rotate +xt neg yt neg translate +dup setflat +/bottom botsv def +0 1 numsteps1 +{ +numsteps1 div taperfcn /frac xdf +bcs 0 +1.0 tint2 tint1 sub frac mul tint1 add sub +put bcs vc +1 index setflat +{ +mark {newpath left bottom right delta rectfill}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +/bottom bottom delta sub def +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/_rfs4 { + /iosv inkoverprint def + /cosv currentoverprint def + /tint2 xdf + /tint1 xdf + bcs exch 1 exch put + /radius xdf + /yt xdf + /xt xdf + tint1 tint2 sub abs + bcs 1 get maxspot + calcgraysteps mul abs round + radius abs adjnumsteps + dup 2 lt {pop 2} if + 1 sub /numsteps1 xdf + radius numsteps1 div 2 div /halfstep xdf + currentflat mark + currentflat clipflatness + { +{ +dup setflat +W +0 1 numsteps1 +{ +dup /radindex xdf +numsteps1 div /frac xdf +bcs 0 +tint2 tint1 sub frac mul tint1 add +put bcs vc +1 index setflat +{ +newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 +{ arc +radindex numsteps1 ne +{ +xt yt +radindex 1 add numsteps1 +div 1 exch sub +radius mul halfstep add +dup xt add yt moveto +360 0 arcn +} if +fill +}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/_rfp4 { + /iosv inkoverprint def + /cosv currentoverprint def + /k2 xdf /y2 xdf /m2 xdf /c2 xdf + /k1 xdf /y1 xdf /m1 xdf /c1 xdf + /radius xdf + /yt xdf + /xt xdf + c1 c2 sub abs + m1 m2 sub abs + y1 y2 sub abs + k1 k2 sub abs + maxcolor + calcgraysteps mul abs round + radius abs adjnumsteps + dup 2 lt {pop 1} if + 1 sub /numsteps1 xdf + radius numsteps1 dup 0 eq {pop} {div} ifelse + 2 div /halfstep xdf + currentflat mark + currentflat clipflatness + { +{ +dup setflat +W +0 1 numsteps1 +{ +dup /radindex xdf +numsteps1 dup 0 eq {pop 0.5 } { div } ifelse +/frac xdf +bc4 0 c2 c1 sub frac mul c1 add put +bc4 1 m2 m1 sub frac mul m1 add put +bc4 2 y2 y1 sub frac mul y1 add put +bc4 3 k2 k1 sub frac mul k1 add put +bc4 vc +1 index setflat +{ +newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 +{ arc +radindex numsteps1 ne +{ +xt yt +radindex 1 add +numsteps1 dup 0 eq {pop} {div} ifelse +1 exch sub +radius mul halfstep add +dup xt add yt moveto +360 0 arcn +} if +fill +}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/lfp4{_lfp4}ndf +/lfs4{_lfs4}ndf +/rfs4{_rfs4}ndf +/rfp4{_rfp4}ndf +/cvc [0 0 0 1] def +/vc{ + AltsysDict /cvc 2 index put + aload length 4 eq + {setcmykcolor} + {setspotcolor} + ifelse +}bdf +/origmtx matrix currentmatrix def +/ImMatrix matrix currentmatrix def +0 setseparationgray +/imgr {1692 1570.1102 2287.2756 2412 } def +/bleed 0 def +/clpr {1692 1570.1102 2287.2756 2412 } def +/xs 1 def +/ys 1 def +/botx 0 def +/overlap 0 def +/wdist 18 def +0 2 mul fhsetspreadsize +0 0 ne {/df 0 def /clipflatness 0 def} if +/maxsteps 256 def +/forcemaxsteps false def +vms +-1845 -1956 translate +/currentpacking defed{false setpacking}if +/spots[ +1 0 0 0 (Process Cyan) false newcmykcustomcolor +0 1 0 0 (Process Magenta) false newcmykcustomcolor +0 0 1 0 (Process Yellow) false newcmykcustomcolor +0 0 0 1 (Process Black) false newcmykcustomcolor +]def +/textopf false def +/curtextmtx{}def +/otw .25 def +/msf{dup/curtextmtx xdf makefont setfont}bdf +/makesetfont/msf load def +/curtextheight{.707104 .707104 curtextmtx dtransform + dup mul exch dup mul add sqrt}bdf +/ta2{ +tempstr 2 index gsave exec grestore +cwidth cheight rmoveto +4 index eq{5 index 5 index rmoveto}if +2 index 2 index rmoveto +}bdf +/ta{exch systemdict/cshow known +{{/cheight xdf/cwidth xdf tempstr 0 2 index put ta2}exch cshow} +{{tempstr 0 2 index put tempstr stringwidth/cheight xdf/cwidth xdf ta2}forall} +ifelse 6{pop}repeat}bdf +/sts{/textopf currentoverprint def vc setoverprint +/ts{awidthshow}def exec textopf setoverprint}bdf +/stol{/xt currentlinewidth def + setlinewidth vc newpath + /ts{{false charpath stroke}ta}def exec + xt setlinewidth}bdf + +/strk{/textopf currentoverprint def vc setoverprint + /ts{{false charpath stroke}ta}def exec + textopf setoverprint + }bdf +n +[] 0 d +3.863708 M +1 w +0 j +0 J +false setoverprint +0 i +false eomode +[0 0 0 1] vc +vms +%white border -- disabled +%1845.2293 2127.8588 m +%2045.9437 2127.8588 L +%2045.9437 1956.1412 L +%1845.2293 1956.1412 L +%1845.2293 2127.8588 L +%0.1417 w +%2 J +%2 M +%[0 0 0 0] vc +%s +n +1950.8 2097.2 m +1958.8 2092.5 1967.3 2089 1975.5 2084.9 C +1976.7 2083.5 1976.1 2081.5 1976.7 2079.9 C +1979.6 2081.1 1981.6 2086.8 1985.3 2084 C +1993.4 2079.3 2001.8 2075.8 2010 2071.7 C +2010.5 2071.5 2010.5 2071.1 2010.8 2070.8 C +2011.2 2064.3 2010.9 2057.5 2011 2050.8 C +2015.8 2046.9 2022.2 2046.2 2026.6 2041.7 C +2026.5 2032.5 2026.8 2022.9 2026.4 2014.1 C +2020.4 2008.3 2015 2002.4 2008.8 1997.1 C +2003.8 1996.8 2000.7 2001.2 1996.1 2002.1 C +1995.2 1996.4 1996.9 1990.5 1995.6 1984.8 C +1989.9 1979 1984.5 1973.9 1978.8 1967.8 C +1977.7 1968.6 1976 1967.6 1974.5 1968.3 C +1967.4 1972.5 1960.1 1976.1 1952.7 1979.3 C +1946.8 1976.3 1943.4 1970.7 1938.5 1966.1 C +1933.9 1966.5 1929.4 1968.8 1925.1 1970.7 C +1917.2 1978.2 1906 1977.9 1897.2 1983.4 C +1893.2 1985.6 1889.4 1988.6 1885 1990.1 C +1884.6 1990.6 1883.9 1991 1883.8 1991.6 C +1883.7 2000.4 1884 2009.9 1883.6 2018.9 C +1887.7 2024 1893.2 2028.8 1898 2033.8 C +1899.1 2035.5 1900.9 2036.8 1902.5 2037.9 C +1903.9 2037.3 1905.2 2036.6 1906.4 2035.5 C +1906.3 2039.7 1906.5 2044.6 1906.1 2048.9 C +1906.3 2049.6 1906.7 2050.2 1907.1 2050.8 C +1913.4 2056 1918.5 2062.7 1924.8 2068.1 C +1926.6 2067.9 1928 2066.9 1929.4 2066 C +1930.2 2071 1927.7 2077.1 1930.6 2081.6 C +1936.6 2086.9 1941.5 2092.9 1947.9 2097.9 C +1949 2098.1 1949.9 2097.5 1950.8 2097.2 C +[0 0 0 0.18] vc +f +0.4 w +S +n +1975.2 2084.7 m +1976.6 2083.4 1975.7 2081.1 1976 2079.4 C +1979.3 2079.5 1980.9 2086.2 1984.8 2084 C +1992.9 2078.9 2001.7 2075.6 2010 2071.2 C +2011 2064.6 2010.2 2057.3 2010.8 2050.6 C +2015.4 2046.9 2021.1 2045.9 2025.9 2042.4 C +2026.5 2033.2 2026.8 2022.9 2025.6 2013.9 C +2020.5 2008.1 2014.5 2003.1 2009.3 1997.6 C +2004.1 1996.7 2000.7 2001.6 1995.9 2002.6 C +1995.2 1996.7 1996.3 1990.2 1994.9 1984.6 C +1989.8 1978.7 1983.6 1973.7 1978.4 1968 C +1977.3 1969.3 1976 1967.6 1974.8 1968.5 C +1967.7 1972.7 1960.4 1976.3 1952.9 1979.6 C +1946.5 1976.9 1943.1 1970.5 1937.8 1966.1 C +1928.3 1968.2 1920.6 1974.8 1911.6 1978.4 C +1901.9 1979.7 1893.9 1986.6 1885 1990.6 C +1884.3 1991 1884.3 1991.7 1884 1992.3 C +1884.5 2001 1884.2 2011 1884.3 2019.9 C +1890.9 2025.3 1895.9 2031.9 1902.3 2037.4 C +1904.2 2037.9 1905.6 2034.2 1906.8 2035.7 C +1907.4 2040.9 1905.7 2046.1 1907.3 2050.8 C +1913.6 2056.2 1919.2 2062.6 1925.1 2067.9 C +1926.9 2067.8 1928 2066.3 1929.6 2065.7 C +1929.9 2070.5 1929.2 2076 1930.1 2080.8 C +1936.5 2086.1 1941.6 2092.8 1948.4 2097.6 C +1957.3 2093.3 1966.2 2088.8 1975.2 2084.7 C +[0 0 0 0] vc +f +S +n +1954.8 2093.8 m +1961.6 2090.5 1968.2 2087 1975 2084 C +1975 2082.8 1975.6 2080.9 1974.8 2080.6 C +1974.3 2075.2 1974.6 2069.6 1974.5 2064 C +1977.5 2059.7 1984.5 2060 1988.9 2056.4 C +1989.5 2055.5 1990.5 2055.3 1990.8 2054.4 C +1991.1 2045.7 1991.4 2036.1 1990.6 2027.8 C +1990.7 2026.6 1992 2027.3 1992.8 2027.1 C +1997 2032.4 2002.6 2037.8 2007.6 2042.2 C +2008.7 2042.3 2007.8 2040.6 2007.4 2040 C +2002.3 2035.6 1997.5 2030 1992.8 2025.2 C +1991.6 2024.7 1990.8 2024.9 1990.1 2025.4 C +1989.4 2024.9 1988.1 2025.2 1987.2 2024.4 C +1987.1 2025.8 1988.3 2026.5 1989.4 2026.8 C +1989.4 2026.6 1989.3 2026.2 1989.6 2026.1 C +1989.9 2026.2 1989.9 2026.6 1989.9 2026.8 C +1989.8 2026.6 1990 2026.5 1990.1 2026.4 C +1990.2 2027 1991.1 2028.3 1990.1 2028 C +1989.9 2037.9 1990.5 2044.1 1989.6 2054.2 C +1985.9 2058 1979.7 2057.4 1976 2061.2 C +1974.5 2061.6 1975.2 2059.9 1974.5 2059.5 C +1973.9 2058 1975.6 2057.8 1975 2056.6 C +1974.5 2057.1 1974.6 2055.3 1973.6 2055.9 C +1971.9 2059.3 1974.7 2062.1 1973.1 2065.5 C +1973.1 2071.2 1972.9 2077 1973.3 2082.5 C +1967.7 2085.6 1962 2088 1956.3 2090.7 C +1953.9 2092.4 1951 2093 1948.6 2094.8 C +1943.7 2089.9 1937.9 2084.3 1933 2079.6 C +1931.3 2076.1 1933.2 2071.3 1932.3 2067.2 C +1931.3 2062.9 1933.3 2060.6 1932 2057.6 C +1932.7 2056.5 1930.9 2053.3 1933.2 2051.8 C +1936.8 2050.1 1940.1 2046.9 1944 2046.8 C +1946.3 2049.7 1949.3 2051.9 1952 2054.4 C +1954.5 2054.2 1956.4 2052.3 1958.7 2051.3 C +1960.8 2050 1963.2 2049 1965.6 2048.4 C +1968.3 2050.8 1970.7 2054.3 1973.6 2055.4 C +1973 2052.2 1969.7 2050.4 1967.6 2048.2 C +1967.1 2046.7 1968.8 2046.6 1969.5 2045.8 C +1972.8 2043.3 1980.6 2043.4 1979.3 2038.4 C +1979.4 2038.6 1979.2 2038.7 1979.1 2038.8 C +1978.7 2038.6 1978.9 2038.1 1978.8 2037.6 C +1978.9 2037.9 1978.7 2038 1978.6 2038.1 C +1978.2 2032.7 1978.4 2027.1 1978.4 2021.6 C +1979.3 2021.1 1980 2020.2 1981.5 2020.1 C +1983.5 2020.5 1984 2021.8 1985.1 2023.5 C +1985.7 2024 1987.4 2023.7 1986 2022.8 C +1984.7 2021.7 1983.3 2020.8 1983.9 2018.7 C +1987.2 2015.9 1993 2015.4 1994.9 2011.5 C +1992.2 2004.9 1999.3 2005.2 2002.1 2002.4 C +2005.9 2002.7 2004.8 1997.4 2009.1 1999 C +2011 1999.3 2010 2002.9 2012.7 2002.4 C +2010.2 2000.7 2009.4 1996.1 2005.5 1998.5 C +2002.1 2000.3 1999 2002.5 1995.4 2003.8 C +1995.2 2003.6 1994.9 2003.3 1994.7 2003.1 C +1994.3 1997 1995.6 1991.1 1994.4 1985.3 C +1994.3 1986 1993.8 1985 1994 1985.6 C +1993.8 1995.4 1994.4 2001.6 1993.5 2011.7 C +1989.7 2015.5 1983.6 2014.9 1979.8 2018.7 C +1978.3 2019.1 1979.1 2017.4 1978.4 2017 C +1977.8 2015.5 1979.4 2015.3 1978.8 2014.1 C +1978.4 2014.6 1978.5 2012.8 1977.4 2013.4 C +1975.8 2016.8 1978.5 2019.6 1976.9 2023 C +1977 2028.7 1976.7 2034.5 1977.2 2040 C +1971.6 2043.1 1965.8 2045.6 1960.1 2048.2 C +1957.7 2049.9 1954.8 2050.5 1952.4 2052.3 C +1947.6 2047.4 1941.8 2041.8 1936.8 2037.2 C +1935.2 2033.6 1937.1 2028.8 1936.1 2024.7 C +1935.1 2020.4 1937.1 2018.1 1935.9 2015.1 C +1936.5 2014.1 1934.7 2010.8 1937.1 2009.3 C +1944.4 2004.8 1952 2000.9 1959.9 1997.8 C +1963.9 1997 1963.9 2001.9 1966.8 2003.3 C +1970.3 2006.9 1973.7 2009.9 1976.9 2012.9 C +1977.9 2013 1977.1 2011.4 1976.7 2010.8 C +1971.6 2006.3 1966.8 2000.7 1962 1995.9 C +1960 1995.2 1960.1 1996.6 1958.2 1995.6 C +1957 1997 1955.1 1998.8 1953.2 1998 C +1951.7 1994.5 1954.1 1993.4 1952.9 1991.1 C +1952.1 1990.5 1953.3 1990.2 1953.2 1989.6 C +1954.2 1986.8 1950.9 1981.4 1954.4 1981.2 C +1954.7 1981.6 1954.7 1981.7 1955.1 1982 C +1961.9 1979.1 1967.6 1975 1974.3 1971.6 C +1974.7 1969.8 1976.7 1969.5 1978.4 1969.7 C +1980.3 1970 1979.3 1973.6 1982 1973.1 C +1975.8 1962.2 1968 1975.8 1960.8 1976.7 C +1956.9 1977.4 1953.3 1982.4 1949.1 1978.8 C +1946 1975.8 1941.2 1971 1939.5 1969.2 C +1938.5 1968.6 1938.9 1967.4 1937.8 1966.8 C +1928.7 1969.4 1920.6 1974.5 1912.4 1979.1 C +1904 1980 1896.6 1985 1889.3 1989.4 C +1887.9 1990.4 1885.1 1990.3 1885 1992.5 C +1885.4 2000.6 1885.2 2012.9 1885.2 2019.9 C +1886.1 2022 1889.7 2019.5 1888.4 2022.8 C +1889 2023.3 1889.8 2024.4 1890.3 2024 C +1891.2 2023.5 1891.8 2028.2 1893.4 2026.6 C +1894.2 2026.3 1893.9 2027.3 1894.4 2027.6 C +1893.4 2027.6 1894.7 2028.3 1894.1 2028.5 C +1894.4 2029.6 1896 2030 1896 2029.2 C +1896.2 2029 1896.3 2029 1896.5 2029.2 C +1896.8 2029.8 1897.3 2030 1897 2030.7 C +1896.5 2030.7 1896.9 2031.5 1897.2 2031.6 C +1898.3 2034 1899.5 2030.6 1899.6 2033.3 C +1898.5 2033 1899.6 2034.4 1900.1 2034.8 C +1901.3 2035.8 1903.2 2034.6 1902.5 2036.7 C +1904.4 2036.9 1906.1 2032.2 1907.6 2035.5 C +1907.5 2040.1 1907.7 2044.9 1907.3 2049.4 C +1908 2050.2 1908.3 2051.4 1909.5 2051.6 C +1910.1 2051.1 1911.6 2051.1 1911.4 2052.3 C +1909.7 2052.8 1912.4 2054 1912.6 2054.7 C +1913.4 2055.2 1913 2053.7 1913.6 2054.4 C +1913.6 2054.5 1913.6 2055.3 1913.6 2054.7 C +1913.7 2054.4 1913.9 2054.4 1914 2054.7 C +1914 2054.9 1914.1 2055.3 1913.8 2055.4 C +1913.7 2056 1915.2 2057.6 1916 2057.6 C +1915.9 2057.3 1916.1 2057.2 1916.2 2057.1 C +1917 2056.8 1916.7 2057.7 1917.2 2058 C +1917 2058.3 1916.7 2058.3 1916.4 2058.3 C +1917.1 2059 1917.3 2060.1 1918.4 2060.4 C +1918.1 2059.2 1919.1 2060.6 1919.1 2059.5 C +1919 2060.6 1920.6 2060.1 1919.8 2061.2 C +1919.6 2061.2 1919.3 2061.2 1919.1 2061.2 C +1919.6 2061.9 1921.4 2064.2 1921.5 2062.6 C +1922.4 2062.1 1921.6 2063.9 1922.2 2064.3 C +1922.9 2067.3 1926.1 2064.3 1925.6 2067.2 C +1927.2 2066.8 1928.4 2064.6 1930.1 2065.2 C +1931.8 2067.8 1931 2071.8 1930.8 2074.8 C +1930.6 2076.4 1930.1 2078.6 1930.6 2080.4 C +1936.6 2085.4 1941.8 2091.6 1948.1 2096.9 C +1950.7 2096.7 1952.6 2094.8 1954.8 2093.8 C +[0 0.33 0.33 0.99] vc +f +S +n +1989.4 2080.6 m +1996.1 2077.3 2002.7 2073.8 2009.6 2070.8 C +2009.6 2069.6 2010.2 2067.7 2009.3 2067.4 C +2008.9 2062 2009.1 2056.4 2009.1 2050.8 C +2012.3 2046.6 2019 2046.6 2023.5 2043.2 C +2024 2042.3 2025.1 2042.1 2025.4 2041.2 C +2025.3 2032.7 2025.6 2023.1 2025.2 2014.6 C +2025 2015.3 2024.6 2014.2 2024.7 2014.8 C +2024.5 2024.7 2025.1 2030.9 2024.2 2041 C +2020.4 2044.8 2014.3 2044.2 2010.5 2048 C +2009 2048.4 2009.8 2046.7 2009.1 2046.3 C +2008.5 2044.8 2010.2 2044.6 2009.6 2043.4 C +2009.1 2043.9 2009.2 2042.1 2008.1 2042.7 C +2006.5 2046.1 2009.3 2048.9 2007.6 2052.3 C +2007.7 2058 2007.5 2063.8 2007.9 2069.3 C +2002.3 2072.4 1996.5 2074.8 1990.8 2077.5 C +1988.4 2079.2 1985.6 2079.8 1983.2 2081.6 C +1980.5 2079 1977.9 2076.5 1975.5 2074.1 C +1975.5 2075.1 1975.5 2076.2 1975.5 2077.2 C +1977.8 2079.3 1980.3 2081.6 1982.7 2083.7 C +1985.3 2083.5 1987.1 2081.6 1989.4 2080.6 C +f +S +n +1930.1 2079.9 m +1931.1 2075.6 1929.2 2071.1 1930.8 2067.2 C +1930.3 2066.3 1930.1 2064.6 1928.7 2065.5 C +1927.7 2066.4 1926.5 2067 1925.3 2067.4 C +1924.5 2066.9 1925.6 2065.7 1924.4 2066 C +1924.2 2067.2 1923.6 2065.5 1923.2 2065.7 C +1922.3 2063.6 1917.8 2062.1 1919.6 2060.4 C +1919.3 2060.5 1919.2 2060.3 1919.1 2060.2 C +1919.7 2060.9 1918.2 2061 1917.6 2060.2 C +1917 2059.6 1916.1 2058.8 1916.4 2058 C +1915.5 2058 1917.4 2057.1 1915.7 2057.8 C +1914.8 2057.1 1913.4 2056.2 1913.3 2054.9 C +1913.1 2055.4 1911.3 2054.3 1910.9 2053.2 C +1910.7 2052.9 1910.2 2052.5 1910.7 2052.3 C +1911.1 2052.5 1910.9 2052 1910.9 2051.8 C +1910.5 2051.2 1909.9 2052.6 1909.2 2051.8 C +1908.2 2051.4 1907.8 2050.2 1907.1 2049.4 C +1907.5 2044.8 1907.3 2040 1907.3 2035.2 C +1905.3 2033 1902.8 2039.3 1902.3 2035.7 C +1899.6 2036 1898.4 2032.5 1896.3 2030.7 C +1895.7 2030.1 1897.5 2030 1896.3 2029.7 C +1896.3 2030.6 1895 2029.7 1894.4 2029.2 C +1892.9 2028.1 1894.2 2027.4 1893.6 2027.1 C +1892.1 2027.9 1891.7 2025.6 1890.8 2024.9 C +1891.1 2024.6 1889.1 2024.3 1888.4 2023 C +1887.5 2022.6 1888.2 2021.9 1888.1 2021.3 C +1886.7 2022 1885.2 2020.4 1884.8 2019.2 C +1884.8 2010 1884.6 2000.2 1885 1991.8 C +1886.9 1989.6 1889.9 1989.3 1892.2 1987.5 C +1898.3 1982.7 1905.6 1980.1 1912.8 1978.6 C +1921 1974.2 1928.8 1968.9 1937.8 1966.6 C +1939.8 1968.3 1938.8 1968.3 1940.4 1970 C +1945.4 1972.5 1947.6 1981.5 1954.6 1979.3 C +1952.3 1981 1950.4 1978.4 1948.6 1977.9 C +1945.1 1973.9 1941.1 1970.6 1938 1966.6 C +1928.4 1968.5 1920.6 1974.8 1911.9 1978.8 C +1907.1 1979.2 1902.6 1981.7 1898.2 1983.6 C +1893.9 1986 1889.9 1989 1885.5 1990.8 C +1884.9 1991.2 1884.8 1991.8 1884.5 1992.3 C +1884.9 2001.3 1884.7 2011.1 1884.8 2019.6 C +1890.6 2025 1896.5 2031.2 1902.3 2036.9 C +1904.6 2037.6 1905 2033 1907.3 2035.5 C +1907.2 2040.2 1907 2044.8 1907.1 2049.6 C +1913.6 2055.3 1918.4 2061.5 1925.1 2067.4 C +1927.3 2068.2 1929.6 2062.5 1930.6 2066.9 C +1929.7 2070.7 1930.3 2076 1930.1 2080.1 C +1935.6 2085.7 1941.9 2090.7 1947.2 2096.7 C +1942.2 2091.1 1935.5 2085.2 1930.1 2079.9 C +[0.18 0.18 0 0.78] vc +f +S +n +1930.8 2061.9 m +1930.3 2057.8 1931.8 2053.4 1931.1 2050.4 C +1931.3 2050.3 1931.7 2050.5 1931.6 2050.1 C +1933 2051.1 1934.4 2049.5 1935.9 2048.7 C +1937 2046.5 1939.5 2047.1 1941.2 2045.1 C +1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C +1934 2039.7 1934.5 2038.1 1933.7 2037.6 C +1934 2033.3 1933.1 2027.9 1934.4 2024.4 C +1934.3 2023.8 1933.9 2022.8 1933 2022.8 C +1931.6 2023.1 1930.5 2024.4 1929.2 2024.9 C +1928.4 2024.5 1929.8 2023.5 1928.7 2023.5 C +1927.7 2024.1 1926.2 2022.6 1925.6 2021.6 C +1926.9 2021.6 1924.8 2020.6 1925.6 2020.4 C +1924.7 2021.7 1923.9 2019.6 1923.2 2019.2 C +1923.3 2018.3 1923.8 2018.1 1923.2 2018 C +1922.9 2017.8 1922.9 2017.5 1922.9 2017.2 C +1922.8 2018.3 1921.3 2017.3 1920.3 2018 C +1916.6 2019.7 1913 2022.1 1910 2024.7 C +1910 2032.9 1910 2041.2 1910 2049.4 C +1915.4 2055.2 1920 2058.7 1925.3 2064.8 C +1927.2 2064 1929 2061.4 1930.8 2061.9 C +[0 0 0 0] vc +f +S +n +1907.6 2030.4 m +1907.5 2027.1 1906.4 2021.7 1908.5 2019.9 C +1908.8 2020.1 1908.9 2019 1909.2 2019.6 C +1910 2019.6 1912 2019.2 1913.1 2018.2 C +1913.7 2016.5 1920.2 2015.7 1917.4 2012.7 C +1918.2 2011.2 1917 2013.8 1917.2 2012 C +1916.9 2012.3 1916 2012.4 1915.2 2012 C +1912.5 2010.5 1916.6 2008.8 1913.6 2009.6 C +1912.6 2009.2 1911.1 2009 1910.9 2007.6 C +1911 1999.2 1911.8 1989.8 1911.2 1982.2 C +1910.1 1981.1 1908.8 1982.2 1907.6 1982.2 C +1900.8 1986.5 1893.2 1988.8 1887.2 1994.2 C +1887.2 2002.4 1887.2 2010.7 1887.2 2018.9 C +1892.6 2024.7 1897.2 2028.2 1902.5 2034.3 C +1904.3 2033.3 1906.2 2032.1 1907.6 2030.4 C +f +S +n +1910.7 2025.4 m +1912.7 2022.4 1916.7 2020.8 1919.8 2018.9 C +1920.2 2018.7 1920.6 2018.6 1921 2018.4 C +1925 2020 1927.4 2028.5 1932 2024.2 C +1932.3 2025 1932.5 2023.7 1932.8 2024.4 C +1932.8 2028 1932.8 2031.5 1932.8 2035 C +1931.9 2033.9 1932.5 2036.3 1932.3 2036.9 C +1933.2 2036.4 1932.5 2038.5 1933 2038.4 C +1933.1 2040.5 1935.6 2042.2 1936.6 2043.2 C +1936.2 2042.4 1935.1 2040.8 1933.7 2040.3 C +1932.2 2034.4 1933.8 2029.8 1933 2023.2 C +1931.1 2024.9 1928.4 2026.4 1926.5 2023.5 C +1925.1 2021.6 1923 2019.8 1921.5 2018.2 C +1917.8 2018.9 1915.2 2022.5 1911.6 2023.5 C +1910.8 2023.8 1911.2 2024.7 1910.4 2025.2 C +1910.9 2031.8 1910.6 2039.1 1910.7 2045.6 C +1910.1 2048 1910.7 2045.9 1911.2 2044.8 C +1910.6 2038.5 1911.2 2031.8 1910.7 2025.4 C +[0.07 0.06 0 0.58] vc +f +S +n +1910.7 2048.9 m +1910.3 2047.4 1911.3 2046.5 1911.6 2045.3 C +1912.9 2045.3 1913.9 2047.1 1915.2 2045.8 C +1915.2 2044.9 1916.6 2043.3 1917.2 2042.9 C +1918.7 2042.9 1919.4 2044.4 1920.5 2043.2 C +1921.2 2042.2 1921.4 2040.9 1922.4 2040.3 C +1924.5 2040.3 1925.7 2040.9 1926.8 2039.6 C +1927.1 2037.9 1926.8 2038.1 1927.7 2037.6 C +1929 2037.5 1930.4 2037 1931.6 2037.2 C +1932.3 2038.2 1933.1 2038.7 1932.8 2040.3 C +1935 2041.8 1935.9 2043.8 1938.5 2044.8 C +1938.6 2045 1938.3 2045.5 1938.8 2045.3 C +1939.1 2042.9 1935.4 2044.2 1935.4 2042.2 C +1932.1 2040.8 1932.8 2037.2 1932 2034.8 C +1932.3 2034 1932.7 2035.4 1932.5 2034.8 C +1931.3 2031.8 1935.5 2020.1 1928.9 2025.9 C +1924.6 2024.7 1922.6 2014.5 1917.4 2020.4 C +1915.5 2022.8 1912 2022.6 1910.9 2025.4 C +1911.5 2031.9 1910.9 2038.8 1911.4 2045.3 C +1911.1 2046.5 1910 2047.4 1910.4 2048.9 C +1915.1 2054.4 1920.4 2058.3 1925.1 2063.8 C +1920.8 2058.6 1914.9 2054.3 1910.7 2048.9 C +[0.4 0.4 0 0] vc +f +S +n +1934.7 2031.9 m +1934.6 2030.7 1934.9 2029.5 1934.4 2028.5 C +1934 2029.5 1934.3 2031.2 1934.2 2032.6 C +1933.8 2031.7 1934.9 2031.6 1934.7 2031.9 C +[0.92 0.92 0 0.67] vc +f +S +n +vmrs +1934.7 2019.4 m +1934.1 2015.3 1935.6 2010.9 1934.9 2007.9 C +1935.1 2007.8 1935.6 2008.1 1935.4 2007.6 C +1936.8 2008.6 1938.2 2007 1939.7 2006.2 C +1940.1 2004.3 1942.7 2005 1943.6 2003.8 C +1945.1 2000.3 1954 2000.8 1950 1996.6 C +1952.1 1993.3 1948.2 1989.2 1951.2 1985.6 C +1953 1981.4 1948.4 1982.3 1947.9 1979.8 C +1945.4 1979.6 1945.1 1975.5 1942.4 1975 C +1942.4 1972.3 1938 1973.6 1938.5 1970.4 C +1937.4 1969 1935.6 1970.1 1934.2 1970.2 C +1927.5 1974.5 1919.8 1976.8 1913.8 1982.2 C +1913.8 1990.4 1913.8 1998.7 1913.8 2006.9 C +1919.3 2012.7 1923.8 2016.2 1929.2 2022.3 C +1931.1 2021.6 1932.8 2018.9 1934.7 2019.4 C +[0 0 0 0] vc +f +0.4 w +2 J +2 M +S +n +2024.2 2038.1 m +2024.1 2029.3 2024.4 2021.7 2024.7 2014.4 C +2024.4 2013.6 2020.6 2013.4 2021.3 2011.2 C +2020.5 2010.3 2018.4 2010.6 2018.9 2008.6 C +2019 2008.8 2018.8 2009 2018.7 2009.1 C +2018.2 2006.7 2015.2 2007.9 2015.3 2005.5 C +2014.7 2004.8 2012.4 2005.1 2013.2 2003.6 C +2012.3 2004.2 2012.8 2002.4 2012.7 2002.6 C +2009.4 2003.3 2011.2 1998.6 2008.4 1999.2 C +2007 1999.1 2006.1 1999.4 2005.7 2000.4 C +2006.9 1998.5 2007.7 2000.5 2009.3 2000.2 C +2009.2 2003.7 2012.4 2002.1 2012.9 2005.2 C +2015.9 2005.6 2015.2 2008.6 2017.7 2008.8 C +2018.4 2009.6 2018.3 2011.4 2019.6 2011 C +2021.1 2011.7 2021.4 2014.8 2023.7 2015.1 C +2023.7 2023.5 2023.9 2031.6 2023.5 2040.5 C +2021.8 2041.7 2020.7 2043.6 2018.4 2043.9 C +2020.8 2042.7 2025.5 2041.8 2024.2 2038.1 C +[0 0.87 0.91 0.83] vc +f +S +n +2023.5 2040 m +2023.5 2031.1 2023.5 2023.4 2023.5 2015.1 C +2020.2 2015 2021.8 2010.3 2018.4 2011 C +2018.6 2007.5 2014.7 2009.3 2014.8 2006.4 C +2011.8 2006.3 2012.2 2002.3 2009.8 2002.4 C +2009.7 2001.5 2009.2 2000.1 2008.4 2000.2 C +2008.7 2000.9 2009.7 2001.2 2009.3 2002.4 C +2008.4 2004.2 2007.5 2003.1 2007.9 2005.5 C +2007.9 2010.8 2007.7 2018.7 2008.1 2023.2 C +2009 2024.3 2007.3 2023.4 2007.9 2024 C +2007.7 2024.6 2007.3 2026.3 2008.6 2027.1 C +2009.7 2026.8 2010 2027.6 2010.5 2028 C +2010.5 2028.2 2010.5 2029.1 2010.5 2028.5 C +2011.5 2028 2010.5 2030 2011.5 2030 C +2014.2 2029.7 2012.9 2032.2 2014.8 2032.6 C +2015.1 2033.6 2015.3 2033 2016 2033.3 C +2017 2033.9 2016.6 2035.4 2017.2 2036.2 C +2018.7 2036.4 2019.2 2039 2021.3 2038.4 C +2021.6 2035.4 2019.7 2029.5 2021.1 2027.3 C +2020.9 2023.5 2021.5 2018.5 2020.6 2016 C +2020.9 2013.9 2021.5 2015.4 2022.3 2014.4 C +2022.2 2015.1 2023.3 2014.8 2023.2 2015.6 C +2022.7 2019.8 2023.3 2024.3 2022.8 2028.5 C +2022.3 2028.2 2022.6 2027.6 2022.5 2027.1 C +2022.5 2027.8 2022.5 2029.2 2022.5 2029.2 C +2022.6 2029.2 2022.7 2029.1 2022.8 2029 C +2023.9 2032.8 2022.6 2037 2023 2040.8 C +2022.3 2041.2 2021.6 2041.5 2021.1 2042.2 C +2022 2041.2 2022.9 2041.4 2023.5 2040 C +[0 1 1 0.23] vc +f +S +n +2009.1 1997.8 m +2003.8 1997.7 2000.1 2002.4 1995.4 2003.1 C +1995 1999.5 1995.2 1995 1995.2 1992 C +1995.2 1995.8 1995 1999.7 1995.4 2003.3 C +2000.3 2002.2 2003.8 1997.9 2009.1 1997.8 C +2012.3 2001.2 2015.6 2004.8 2018.7 2008.1 C +2021.6 2011.2 2027.5 2013.9 2025.9 2019.9 C +2026.1 2017.9 2025.6 2016.2 2025.4 2014.4 C +2020.2 2008.4 2014 2003.6 2009.1 1997.8 C +[0.18 0.18 0 0.78] vc +f +S +n +2009.3 1997.8 m +2008.7 1997.4 2007.9 1997.6 2007.2 1997.6 C +2007.9 1997.6 2008.9 1997.4 2009.6 1997.8 C +2014.7 2003.6 2020.8 2008.8 2025.9 2014.8 C +2025.8 2017.7 2026.1 2014.8 2025.6 2014.1 C +2020.4 2008.8 2014.8 2003.3 2009.3 1997.8 C +[0.07 0.06 0 0.58] vc +f +S +n +2009.6 1997.6 m +2009 1997.1 2008.1 1997.4 2007.4 1997.3 C +2008.1 1997.4 2009 1997.1 2009.6 1997.6 C +2014.8 2003.7 2021.1 2008.3 2025.9 2014.4 C +2021.1 2008.3 2014.7 2003.5 2009.6 1997.6 C +[0.4 0.4 0 0] vc +f +S +n +2021.8 2011.5 m +2021.9 2012.2 2022.3 2013.5 2023.7 2013.6 C +2023.4 2012.7 2022.8 2011.8 2021.8 2011.5 C +[0 0.33 0.33 0.99] vc +f +S +n +2021.1 2042 m +2022.1 2041.1 2020.9 2040.2 2020.6 2039.6 C +2018.4 2039.5 2018.1 2036.9 2016.3 2036.4 C +2015.8 2035.5 2015.3 2033.8 2014.8 2033.6 C +2012.4 2033.8 2013 2030.4 2010.5 2030.2 C +2009.6 2028.9 2009.6 2028.3 2008.4 2028 C +2006.9 2026.7 2007.5 2024.3 2006 2023.2 C +2006.6 2023.2 2005.7 2023.3 2005.7 2023 C +2006.4 2022.5 2006.3 2021.1 2006.7 2020.6 C +2006.6 2015 2006.9 2009 2006.4 2003.8 C +2006.9 2002.5 2007.6 2001.1 2006.9 2000.7 C +2004.6 2003.6 2003 2002.9 2000.2 2004.3 C +1999.3 2005.8 1997.9 2006.3 1996.1 2006.7 C +1995.7 2008.9 1996 2011.1 1995.9 2012.9 C +1993.4 2015.1 1990.5 2016.2 1987.7 2017.7 C +1987.1 2019.3 1991.1 2019.4 1990.4 2021.3 C +1990.5 2021.5 1991.9 2022.3 1992 2023 C +1994.8 2024.4 1996.2 2027.5 1998.5 2030 C +2002.4 2033 2005.2 2037.2 2008.8 2041 C +2010.2 2041.3 2011.6 2042 2011 2043.9 C +2011.2 2044.8 2010.1 2045.3 2010.5 2046.3 C +2013.8 2044.8 2017.5 2043.4 2021.1 2042 C +[0 0.5 0.5 0.2] vc +f +S +n +2019.4 2008.8 m +2018.9 2009.2 2019.3 2009.9 2019.6 2010.3 C +2022.2 2011.5 2020.3 2009.1 2019.4 2008.8 C +[0 0.33 0.33 0.99] vc +f +S +n +2018 2007.4 m +2015.7 2006.7 2015.3 2003.6 2012.9 2002.8 C +2013.5 2003.7 2013.5 2005.1 2015.6 2005.2 C +2016.4 2006.1 2015.7 2007.7 2018 2007.4 C +f +S +n +vmrs +1993.5 2008.8 m +1993.4 2000 1993.7 1992.5 1994 1985.1 C +1993.7 1984.3 1989.9 1984.1 1990.6 1982 C +1989.8 1981.1 1987.7 1981.4 1988.2 1979.3 C +1988.3 1979.6 1988.1 1979.7 1988 1979.8 C +1987.5 1977.5 1984.5 1978.6 1984.6 1976.2 C +1983.9 1975.5 1981.7 1975.8 1982.4 1974.3 C +1981.6 1974.9 1982.1 1973.1 1982 1973.3 C +1979 1973.7 1980 1968.8 1976.9 1969.7 C +1975.9 1969.8 1975.3 1970.3 1975 1971.2 C +1976.2 1969.2 1977 1971.2 1978.6 1970.9 C +1978.5 1974.4 1981.7 1972.8 1982.2 1976 C +1985.2 1976.3 1984.5 1979.3 1987 1979.6 C +1987.7 1980.3 1987.5 1982.1 1988.9 1981.7 C +1990.4 1982.4 1990.7 1985.5 1993 1985.8 C +1992.9 1994.3 1993.2 2002.3 1992.8 2011.2 C +1991.1 2012.4 1990 2014.4 1987.7 2014.6 C +1990.1 2013.4 1994.7 2012.6 1993.5 2008.8 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1992.8 2010.8 m +1992.8 2001.8 1992.8 1994.1 1992.8 1985.8 C +1989.5 1985.7 1991.1 1981.1 1987.7 1981.7 C +1987.9 1978.2 1983.9 1980 1984.1 1977.2 C +1981.1 1977 1981.5 1973 1979.1 1973.1 C +1979 1972.2 1978.5 1970.9 1977.6 1970.9 C +1977.9 1971.6 1979 1971.9 1978.6 1973.1 C +1977.6 1974.9 1976.8 1973.9 1977.2 1976.2 C +1977.2 1981.5 1977 1989.4 1977.4 1994 C +1978.3 1995 1976.6 1994.1 1977.2 1994.7 C +1977 1995.3 1976.6 1997 1977.9 1997.8 C +1979 1997.5 1979.3 1998.3 1979.8 1998.8 C +1979.8 1998.9 1979.8 1999.8 1979.8 1999.2 C +1980.8 1998.7 1979.7 2000.7 1980.8 2000.7 C +1983.5 2000.4 1982.1 2003 1984.1 2003.3 C +1984.4 2004.3 1984.5 2003.7 1985.3 2004 C +1986.3 2004.6 1985.9 2006.1 1986.5 2006.9 C +1988 2007.1 1988.4 2009.7 1990.6 2009.1 C +1990.9 2006.1 1989 2000.2 1990.4 1998 C +1990.2 1994.3 1990.8 1989.2 1989.9 1986.8 C +1990.2 1984.7 1990.8 1986.2 1991.6 1985.1 C +1991.5 1985.9 1992.6 1985.5 1992.5 1986.3 C +1992 1990.5 1992.6 1995 1992 1999.2 C +1991.6 1998.9 1991.9 1998.3 1991.8 1997.8 C +1991.8 1998.5 1991.8 2000 1991.8 2000 C +1991.9 1999.9 1992 1999.8 1992 1999.7 C +1993.2 2003.5 1991.9 2007.7 1992.3 2011.5 C +1991.6 2012 1990.9 2012.2 1990.4 2012.9 C +1991.3 2011.9 1992.2 2012.1 1992.8 2010.8 C +[0 1 1 0.23] vc +f +S +n +1978.4 1968.5 m +1977 1969.2 1975.8 1968.2 1974.5 1969 C +1968.3 1973 1961.6 1976 1955.1 1979.1 C +1962 1975.9 1968.8 1972.5 1975.5 1968.8 C +1976.5 1968.8 1977.6 1968.8 1978.6 1968.8 C +1981.7 1972.1 1984.8 1975.7 1988 1978.8 C +1990.9 1981.9 1996.8 1984.6 1995.2 1990.6 C +1995.3 1988.6 1994.9 1986.9 1994.7 1985.1 C +1989.5 1979.1 1983.3 1974.3 1978.4 1968.5 C +[0.18 0.18 0 0.78] vc +f +S +n +1978.4 1968.3 m +1977.9 1968.7 1977.1 1968.5 1976.4 1968.5 C +1977.3 1968.8 1978.1 1967.9 1978.8 1968.5 C +1984 1974.3 1990.1 1979.5 1995.2 1985.6 C +1995.1 1988.4 1995.3 1985.6 1994.9 1984.8 C +1989.5 1979.4 1983.9 1973.8 1978.4 1968.3 C +[0.07 0.06 0 0.58] vc +f +S +n +1978.6 1968 m +1977.9 1968 1977.4 1968.6 1978.4 1968 C +1983.9 1973.9 1990.1 1979.1 1995.2 1985.1 C +1990.2 1979 1983.8 1974.1 1978.6 1968 C +[0.4 0.4 0 0] vc +f +S +n +1991.1 1982.2 m +1991.2 1982.9 1991.6 1984.2 1993 1984.4 C +1992.6 1983.5 1992.1 1982.5 1991.1 1982.2 C +[0 0.33 0.33 0.99] vc +f +S +n +1990.4 2012.7 m +1991.4 2011.8 1990.2 2010.9 1989.9 2010.3 C +1987.7 2010.2 1987.4 2007.6 1985.6 2007.2 C +1985.1 2006.2 1984.6 2004.5 1984.1 2004.3 C +1981.7 2004.5 1982.3 2001.2 1979.8 2000.9 C +1978.8 1999.6 1978.8 1999.1 1977.6 1998.8 C +1976.1 1997.4 1976.7 1995 1975.2 1994 C +1975.8 1994 1975 1994 1975 1993.7 C +1975.7 1993.2 1975.6 1991.8 1976 1991.3 C +1975.9 1985.7 1976.1 1979.7 1975.7 1974.5 C +1976.2 1973.3 1976.9 1971.8 1976.2 1971.4 C +1973.9 1974.3 1972.2 1973.6 1969.5 1975 C +1967.9 1977.5 1963.8 1977.1 1961.8 1980 C +1959 1980 1957.6 1983 1954.8 1982.9 C +1953.8 1984.2 1954.8 1985.7 1955.1 1987.2 C +1956.2 1989.5 1959.7 1990.1 1959.9 1991.8 C +1965.9 1998 1971.8 2005.2 1978.1 2011.7 C +1979.5 2012 1980.9 2012.7 1980.3 2014.6 C +1980.5 2015.6 1979.4 2016 1979.8 2017 C +1983 2015.6 1986.8 2014.1 1990.4 2012.7 C +[0 0.5 0.5 0.2] vc +f +S +n +1988.7 1979.6 m +1988.2 1979.9 1988.6 1980.6 1988.9 1981 C +1991.4 1982.2 1989.6 1979.9 1988.7 1979.6 C +[0 0.33 0.33 0.99] vc +f +S +n +1987.2 1978.1 m +1985 1977.5 1984.6 1974.3 1982.2 1973.6 C +1982.7 1974.5 1982.8 1975.8 1984.8 1976 C +1985.7 1976.9 1985 1978.4 1987.2 1978.1 C +f +S +n +1975.5 2084 m +1975.5 2082 1975.3 2080 1975.7 2078.2 C +1978.8 2079 1980.9 2085.5 1984.8 2083.5 C +1993 2078.7 2001.6 2075 2010 2070.8 C +2010.1 2064 2009.9 2057.2 2010.3 2050.6 C +2014.8 2046.2 2020.9 2045.7 2025.6 2042 C +2026.1 2035.1 2025.8 2028 2025.9 2021.1 C +2025.8 2027.8 2026.1 2034.6 2025.6 2041.2 C +2022.2 2044.9 2017.6 2046.8 2012.9 2048 C +2012.5 2049.5 2010.4 2049.4 2009.8 2051.1 C +2009.9 2057.6 2009.6 2064.2 2010 2070.5 C +2001.2 2075.4 1992 2079.1 1983.2 2084 C +1980.3 2082.3 1977.8 2079.2 1975.2 2077.5 C +1974.9 2079.9 1977.2 2084.6 1973.3 2085.2 C +1964.7 2088.6 1956.8 2093.7 1948.1 2097.2 C +1949 2097.3 1949.6 2096.9 1950.3 2096.7 C +1958.4 2091.9 1967.1 2088.2 1975.5 2084 C +[0.18 0.18 0 0.78] vc +f +S +n +vmrs +1948.6 2094.5 m +1950.2 2093.7 1951.8 2092.9 1953.4 2092.1 C +1951.8 2092.9 1950.2 2093.7 1948.6 2094.5 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1971.6 2082.3 m +1971.6 2081.9 1970.7 2081.1 1970.9 2081.3 C +1970.7 2081.6 1970.6 2081.6 1970.4 2081.3 C +1970.8 2080.1 1968.7 2081.7 1968.3 2080.8 C +1966.6 2080.9 1966.7 2078 1964.2 2078.2 C +1964.8 2075 1960.1 2075.8 1960.1 2072.9 C +1958 2072.3 1957.5 2069.3 1955.3 2069.3 C +1953.9 2070.9 1948.8 2067.8 1950 2072 C +1949 2074 1943.2 2070.6 1944 2074.8 C +1942.2 2076.6 1937.6 2073.9 1938 2078.2 C +1936.7 2078.6 1935 2078.6 1933.7 2078.2 C +1933.5 2080 1936.8 2080.7 1937.3 2082.8 C +1939.9 2083.5 1940.6 2086.4 1942.6 2088 C +1945.2 2089.2 1946 2091.3 1948.4 2093.6 C +1956 2089.5 1963.9 2086.1 1971.6 2082.3 C +[0 0.01 1 0] vc +f +S +n +1958.2 2089.7 m +1956.4 2090 1955.6 2091.3 1953.9 2091.9 C +1955.6 2091.9 1956.5 2089.7 1958.2 2089.7 C +[0 0.87 0.91 0.83] vc +f +S +n +1929.9 2080.4 m +1929.5 2077.3 1929.7 2073.9 1929.6 2070.8 C +1929.8 2074.1 1929.2 2077.8 1930.1 2080.8 C +1935.8 2085.9 1941.4 2091.3 1946.9 2096.9 C +1941.2 2091 1935.7 2086 1929.9 2080.4 C +[0.4 0.4 0 0] vc +f +S +n +1930.1 2080.4 m +1935.8 2086 1941.5 2090.7 1946.9 2096.7 C +1941.5 2090.9 1935.7 2085.8 1930.1 2080.4 C +[0.07 0.06 0 0.58] vc +f +S +n +1940.9 2087.1 m +1941.7 2088 1944.8 2090.6 1943.6 2089.2 C +1942.5 2089 1941.6 2087.7 1940.9 2087.1 C +[0 0.87 0.91 0.83] vc +f +S +n +1972.8 2082.8 m +1973 2075.3 1972.4 2066.9 1973.3 2059.5 C +1972.5 2058.9 1972.8 2057.3 1973.1 2056.4 C +1974.8 2055.2 1973.4 2055.5 1972.4 2055.4 C +1970.1 2053.2 1967.9 2050.9 1965.6 2048.7 C +1960.9 2049.9 1956.9 2052.7 1952.4 2054.7 C +1949.3 2052.5 1946.3 2049.5 1943.6 2046.8 C +1939.9 2047.7 1936.8 2050.1 1933.5 2051.8 C +1930.9 2054.9 1933.5 2056.2 1932.3 2059.7 C +1933.2 2059.7 1932.2 2060.5 1932.5 2060.2 C +1933.2 2062.5 1931.6 2064.6 1932.5 2067.4 C +1932.9 2069.7 1932.7 2072.2 1932.8 2074.6 C +1933.6 2070.6 1932.2 2066.3 1933 2062.6 C +1934.4 2058.2 1929.8 2053.5 1935.2 2051.1 C +1937.7 2049.7 1940.2 2048 1942.8 2046.8 C +1945.9 2049.2 1948.8 2052 1951.7 2054.7 C +1952.7 2054.7 1953.6 2054.6 1954.4 2054.2 C +1958.1 2052.5 1961.7 2049.3 1965.9 2049.2 C +1968.2 2052.8 1975.2 2055 1972.6 2060.9 C +1973.3 2062.4 1972.2 2065.2 1972.6 2067.6 C +1972.7 2072.6 1972.4 2077.7 1972.8 2082.5 C +1968.1 2084.9 1963.5 2087.5 1958.7 2089.5 C +1963.5 2087.4 1968.2 2085 1972.8 2082.8 C +f +S +n +1935.2 2081.1 m +1936.8 2083.4 1938.6 2084.6 1940.4 2086.6 C +1938.8 2084.4 1936.7 2083.4 1935.2 2081.1 C +f +S +n +1983.2 2081.3 m +1984.8 2080.5 1986.3 2079.7 1988 2078.9 C +1986.3 2079.7 1984.8 2080.5 1983.2 2081.3 C +f +S +n +2006.2 2069.1 m +2006.2 2068.7 2005.2 2067.9 2005.5 2068.1 C +2005.3 2068.4 2005.2 2068.4 2005 2068.1 C +2005.4 2066.9 2003.3 2068.5 2002.8 2067.6 C +2001.2 2067.7 2001.2 2064.8 1998.8 2065 C +1999.4 2061.8 1994.7 2062.6 1994.7 2059.7 C +1992.4 2059.5 1992.4 2055.8 1990.1 2056.8 C +1985.9 2059.5 1981.1 2061 1976.9 2063.8 C +1977.2 2067.6 1974.9 2074.2 1978.8 2075.8 C +1979.6 2077.8 1981.7 2078.4 1982.9 2080.4 C +1990.6 2076.3 1998.5 2072.9 2006.2 2069.1 C +[0 0.01 1 0] vc +f +S +n +vmrs +1992.8 2076.5 m +1991 2076.8 1990.2 2078.1 1988.4 2078.7 C +1990.2 2078.7 1991 2076.5 1992.8 2076.5 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1975.5 2073.4 m +1976.1 2069.7 1973.9 2064.6 1977.4 2062.4 C +1973.9 2064.5 1976.1 2069.9 1975.5 2073.6 C +1976 2074.8 1979.3 2077.4 1978.1 2076 C +1977 2075.7 1975.8 2074.5 1975.5 2073.4 C +f +S +n +2007.4 2069.6 m +2007.6 2062.1 2007 2053.7 2007.9 2046.3 C +2007.1 2045.7 2007.3 2044.1 2007.6 2043.2 C +2009.4 2042 2007.9 2042.3 2006.9 2042.2 C +2002.2 2037.4 1996.7 2032.4 1992.5 2027.3 C +1992 2027.3 1991.6 2027.3 1991.1 2027.3 C +1991.4 2035.6 1991.4 2045.6 1991.1 2054.4 C +1990.5 2055.5 1988.4 2056.6 1990.6 2055.4 C +1991.6 2055.4 1991.6 2054.1 1991.6 2053.2 C +1990.8 2044.7 1991.9 2035.4 1991.6 2027.6 C +1991.8 2027.6 1992 2027.6 1992.3 2027.6 C +1997 2032.8 2002.5 2037.7 2007.2 2042.9 C +2007.3 2044.8 2006.7 2047.4 2007.6 2048.4 C +2006.9 2055.1 2007.1 2062.5 2007.4 2069.3 C +2002.7 2071.7 1998.1 2074.3 1993.2 2076.3 C +1998 2074.2 2002.7 2071.8 2007.4 2069.6 C +f +S +n +2006.7 2069.1 m +2006.3 2068.6 2005.9 2067.7 2005.7 2066.9 C +2005.7 2059.7 2005.9 2051.4 2005.5 2045.1 C +2004.9 2045.3 2004.7 2044.5 2004.3 2045.3 C +2005.1 2045.3 2004.2 2045.8 2004.8 2046 C +2004.8 2052.2 2004.8 2059.2 2004.8 2064.5 C +2005.7 2065.7 2005.1 2065.7 2005 2066.7 C +2003.8 2067 2002.7 2067.2 2001.9 2066.4 C +2001.3 2064.6 1998 2063.1 1998 2061.9 C +1996.1 2062.3 1996.6 2058.3 1994.2 2058.8 C +1992.6 2057.7 1992.7 2054.8 1989.9 2056.6 C +1985.6 2059.3 1980.9 2060.8 1976.7 2063.6 C +1976 2066.9 1976 2071.2 1976.7 2074.6 C +1977.6 2070.8 1973.1 2062.1 1980.5 2061.2 C +1984.3 2060.3 1987.5 2058.2 1990.8 2056.4 C +1991.7 2056.8 1992.9 2057.2 1993.5 2059.2 C +1994.3 2058.6 1994.4 2060.6 1994.7 2059.2 C +1995.3 2062.7 1999.2 2061.4 1998.8 2064.8 C +2001.8 2065.4 2002.5 2068.4 2005.2 2067.4 C +2004.9 2067.9 2006 2068 2006.4 2069.1 C +2001.8 2071.1 1997.4 2073.9 1992.8 2075.8 C +1997.5 2073.8 2002 2071.2 2006.7 2069.1 C +[0 0.2 1 0] vc +f +S +n +1988.7 2056.6 m +1985.1 2058.7 1981.1 2060.1 1977.6 2061.9 C +1981.3 2060.5 1985.6 2058.1 1988.7 2056.6 C +[0 0.87 0.91 0.83] vc +f +S +n +1977.9 2059.5 m +1975.7 2064.5 1973.7 2054.7 1975.2 2060.9 C +1976 2060.6 1977.6 2059.7 1977.9 2059.5 C +f +S +n +1989.6 2051.3 m +1990.1 2042.3 1989.8 2036.6 1989.9 2028 C +1989.8 2027 1990.8 2028.3 1990.1 2027.3 C +1988.9 2026.7 1986.7 2026.9 1986.8 2024.7 C +1987.4 2023 1985.9 2024.6 1985.1 2023.7 C +1984.1 2021.4 1982.5 2020.5 1980.3 2020.6 C +1979.9 2020.8 1979.5 2021.1 1979.3 2021.6 C +1979.7 2025.8 1978.4 2033 1979.6 2038.1 C +1983.7 2042.9 1968.8 2044.6 1978.8 2042.7 C +1979.3 2042.3 1979.6 2041.9 1980 2041.5 C +1980 2034.8 1980 2027 1980 2021.6 C +1981.3 2020.5 1981.7 2021.5 1982.9 2021.8 C +1983.6 2024.7 1986.1 2023.8 1986.8 2026.4 C +1987.1 2027.7 1988.6 2027.1 1989.2 2028.3 C +1989.1 2036.7 1989.3 2044.8 1988.9 2053.7 C +1987.2 2054.9 1986.2 2056.8 1983.9 2057.1 C +1986.3 2055.9 1990.9 2055 1989.6 2051.3 C +f +S +n +1971.6 2078.9 m +1971.4 2070.5 1972.1 2062.2 1971.6 2055.9 C +1969.9 2053.7 1967.6 2051.7 1965.6 2049.6 C +1961.4 2050.4 1957.6 2053.6 1953.4 2055.2 C +1949.8 2055.6 1948.2 2051.2 1945.5 2049.6 C +1945.1 2048.8 1944.5 2047.9 1943.6 2047.5 C +1940.1 2047.8 1937.3 2051 1934 2052.3 C +1933.7 2052.6 1933.7 2053 1933.2 2053.2 C +1933.7 2060.8 1933.4 2067.2 1933.5 2074.6 C +1933.8 2068.1 1934 2060.9 1933.2 2054 C +1935.3 2050.9 1939.3 2049.6 1942.4 2047.5 C +1942.8 2047.5 1943.4 2047.4 1943.8 2047.7 C +1947.1 2050.2 1950.3 2057.9 1955.3 2054.4 C +1955.4 2054.4 1955.5 2054.3 1955.6 2054.2 C +1955.9 2057.6 1956.1 2061.8 1955.3 2064.8 C +1955.4 2064.3 1955.1 2063.8 1955.6 2063.6 C +1956 2066.6 1955.3 2068.7 1958.7 2069.8 C +1959.2 2071.7 1961.4 2071.7 1962 2074.1 C +1964.4 2074.2 1964 2077.7 1967.3 2078.4 C +1967 2079.7 1968.1 2079.9 1969 2080.1 C +1971.1 2079.9 1970 2079.2 1970.4 2078 C +1969.5 2077.2 1970.3 2075.9 1969.7 2075.1 C +1970.1 2069.8 1970.1 2063.6 1969.7 2058.8 C +1969.2 2058.5 1970 2058.1 1970.2 2057.8 C +1970.4 2058.3 1971.2 2057.7 1971.4 2058.3 C +1971.5 2065.3 1971.2 2073.6 1971.6 2081.1 C +1974.1 2081.4 1969.8 2084.3 1972.4 2082.5 C +1971.9 2081.4 1971.6 2080.2 1971.6 2078.9 C +[0 0.4 1 0] vc +f +S +n +1952.4 2052 m +1954.1 2051.3 1955.6 2050.4 1957.2 2049.6 C +1955.6 2050.4 1954.1 2051.3 1952.4 2052 C +[0 0.87 0.91 0.83] vc +f +S +n +1975.5 2039.8 m +1975.5 2039.4 1974.5 2038.7 1974.8 2038.8 C +1974.6 2039.1 1974.5 2039.1 1974.3 2038.8 C +1974.6 2037.6 1972.5 2039.3 1972.1 2038.4 C +1970.4 2038.4 1970.5 2035.5 1968 2035.7 C +1968.6 2032.5 1964 2033.3 1964 2030.4 C +1961.9 2029.8 1961.4 2026.8 1959.2 2026.8 C +1957.7 2028.5 1952.6 2025.3 1953.9 2029.5 C +1952.9 2031.5 1947 2028.2 1947.9 2032.4 C +1946 2034.2 1941.5 2031.5 1941.9 2035.7 C +1940.6 2036.1 1938.9 2036.1 1937.6 2035.7 C +1937.3 2037.5 1940.7 2038.2 1941.2 2040.3 C +1943.7 2041.1 1944.4 2043.9 1946.4 2045.6 C +1949.1 2046.7 1949.9 2048.8 1952.2 2051.1 C +1959.9 2047.1 1967.7 2043.6 1975.5 2039.8 C +[0 0.01 1 0] vc +f +S +n +vmrs +1962 2047.2 m +1960.2 2047.5 1959.5 2048.9 1957.7 2049.4 C +1959.5 2049.5 1960.3 2047.2 1962 2047.2 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +2012.4 2046.3 m +2010.3 2051.3 2008.3 2041.5 2009.8 2047.7 C +2010.5 2047.4 2012.2 2046.5 2012.4 2046.3 C +f +S +n +1944.8 2044.6 m +1945.5 2045.6 1948.6 2048.1 1947.4 2046.8 C +1946.3 2046.5 1945.5 2045.2 1944.8 2044.6 C +f +S +n +1987.2 2054.9 m +1983.7 2057.3 1979.6 2058 1976 2060.2 C +1974.7 2058.2 1977.2 2055.8 1974.3 2054.9 C +1973.1 2052 1970.4 2050.2 1968 2048 C +1968 2047.7 1968 2047.4 1968.3 2047.2 C +1969.5 2046.1 1983 2040.8 1972.4 2044.8 C +1971.2 2046.6 1967.9 2046 1968 2048.2 C +1970.5 2050.7 1973.8 2052.6 1974.3 2055.6 C +1975.1 2055 1975.7 2056.7 1975.7 2057.1 C +1975.7 2058.2 1974.8 2059.3 1975.5 2060.4 C +1979.3 2058.2 1983.9 2057.7 1987.2 2054.9 C +[0.18 0.18 0 0.78] vc +f +S +n +1967.8 2047.5 m +1968.5 2047 1969.1 2046.5 1969.7 2046 C +1969.1 2046.5 1968.5 2047 1967.8 2047.5 C +[0 0.87 0.91 0.83] vc +f +S +n +1976.7 2040.3 m +1976.9 2032.8 1976.3 2024.4 1977.2 2017 C +1976.4 2016.5 1976.6 2014.8 1976.9 2013.9 C +1978.7 2012.7 1977.2 2013 1976.2 2012.9 C +1971.5 2008.1 1965.9 2003.1 1961.8 1998 C +1960.9 1998 1960.1 1998 1959.2 1998 C +1951.5 2001.1 1944.3 2005.5 1937.1 2009.6 C +1935 2012.9 1937 2013.6 1936.1 2017.2 C +1937.1 2017.2 1936 2018 1936.4 2017.7 C +1937 2020.1 1935.5 2022.1 1936.4 2024.9 C +1936.8 2027.2 1936.5 2029.7 1936.6 2032.1 C +1937.4 2028.2 1936 2023.8 1936.8 2020.1 C +1938.3 2015.7 1933.6 2011 1939 2008.6 C +1945.9 2004.5 1953.1 2000.3 1960.6 1998.3 C +1960.9 1998.3 1961.3 1998.3 1961.6 1998.3 C +1966.2 2003.5 1971.8 2008.4 1976.4 2013.6 C +1976.6 2015.5 1976 2018.1 1976.9 2019.2 C +1976.1 2025.8 1976.4 2033.2 1976.7 2040 C +1971.9 2042.4 1967.4 2045 1962.5 2047 C +1967.3 2044.9 1972 2042.6 1976.7 2040.3 C +f +S +n +1939 2038.6 m +1940.6 2040.9 1942.5 2042.1 1944.3 2044.1 C +1942.7 2041.9 1940.6 2040.9 1939 2038.6 C +f +S +n +2006.2 2065.7 m +2006 2057.3 2006.7 2049 2006.2 2042.7 C +2002.1 2038.4 1997.7 2033.4 1993 2030 C +1992.9 2029.3 1992.5 2028.6 1992 2028.3 C +1992.1 2036.6 1991.9 2046.2 1992.3 2054.9 C +1990.8 2056.2 1989 2056.7 1987.5 2058 C +1988.7 2057.7 1990.7 2054.4 1993 2056.4 C +1993.4 2058.8 1996 2058.2 1996.6 2060.9 C +1999 2061 1998.5 2064.5 2001.9 2065.2 C +2001.5 2066.5 2002.7 2066.7 2003.6 2066.9 C +2005.7 2066.7 2004.6 2066 2005 2064.8 C +2004 2064 2004.8 2062.7 2004.3 2061.9 C +2004.6 2056.6 2004.6 2050.4 2004.3 2045.6 C +2003.7 2045.3 2004.6 2044.9 2004.8 2044.6 C +2005 2045.1 2005.7 2044.5 2006 2045.1 C +2006 2052.1 2005.8 2060.4 2006.2 2067.9 C +2008.7 2068.2 2004.4 2071.1 2006.9 2069.3 C +2006.4 2068.2 2006.2 2067 2006.2 2065.7 C +[0 0.4 1 0] vc +f +S +n +2021.8 2041.7 m +2018.3 2044.1 2014.1 2044.8 2010.5 2047 C +2009.3 2045 2011.7 2042.6 2008.8 2041.7 C +2004.3 2035.1 1997.6 2030.9 1993 2024.4 C +1992.1 2024 1991.5 2024.3 1990.8 2024 C +1993.2 2023.9 1995.3 2027.1 1996.8 2029 C +2000.4 2032.6 2004.9 2036.9 2008.4 2040.8 C +2008.2 2043.1 2011.4 2042.8 2009.8 2045.8 C +2009.8 2046.3 2009.7 2046.9 2010 2047.2 C +2013.8 2045 2018.5 2044.5 2021.8 2041.7 C +[0.18 0.18 0 0.78] vc +f +S +n +2001.6 2034 m +2000.7 2033.1 1999.9 2032.3 1999 2031.4 C +1999.9 2032.3 2000.7 2033.1 2001.6 2034 C +[0 0.87 0.91 0.83] vc +f +S +n +vmrs +1989.4 2024.4 m +1989.5 2025.4 1988.6 2024.3 1988.9 2024.7 C +1990.5 2025.8 1990.7 2024.2 1992.8 2024.9 C +1993.8 2025.9 1995 2027.1 1995.9 2028 C +1994.3 2026 1991.9 2023.4 1989.4 2024.4 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1984.8 2019.9 m +1984.6 2018.6 1986.3 2017.2 1987.7 2016.8 C +1987.2 2017.5 1982.9 2017.9 1984.4 2020.6 C +1984.1 2019.9 1984.9 2020 1984.8 2019.9 C +f +S +n +1981.7 2017 m +1979.6 2022 1977.6 2012.3 1979.1 2018.4 C +1979.8 2018.1 1981.5 2017.2 1981.7 2017 C +f +S +n +1884.3 2019.2 m +1884.7 2010.5 1884.5 2000.6 1884.5 1991.8 C +1886.6 1989.3 1889.9 1988.9 1892.4 1987 C +1890.8 1988.7 1886 1989.1 1884.3 1992.3 C +1884.7 2001 1884.5 2011.3 1884.5 2019.9 C +1891 2025.1 1895.7 2031.5 1902 2036.9 C +1896.1 2031 1890 2024.9 1884.3 2019.2 C +[0.07 0.06 0 0.58] vc +f +S +n +1884 2019.4 m +1884.5 2010.6 1884.2 2000.4 1884.3 1991.8 C +1884.8 1990.4 1887.8 1989 1884.8 1990.8 C +1884.3 1991.3 1884.3 1992 1884 1992.5 C +1884.5 2001.2 1884.2 2011.1 1884.3 2019.9 C +1887.9 2023.1 1891.1 2026.4 1894.4 2030 C +1891.7 2026.1 1887.1 2022.9 1884 2019.4 C +[0.4 0.4 0 0] vc +f +S +n +1885 2011.7 m +1885 2006.9 1885 2001.9 1885 1997.1 C +1885 2001.9 1885 2006.9 1885 2011.7 C +[0 0.87 0.91 0.83] vc +f +S +n +1975.5 2036.4 m +1975.2 2028 1976 2019.7 1975.5 2013.4 C +1971.1 2008.5 1965.6 2003.6 1961.6 1999 C +1958.8 1998 1956 2000 1953.6 2001.2 C +1948.2 2004.7 1941.9 2006.5 1937.1 2010.8 C +1937.5 2018.3 1937.3 2024.7 1937.3 2032.1 C +1937.6 2025.6 1937.9 2018.4 1937.1 2011.5 C +1937.3 2011 1937.6 2010.5 1937.8 2010 C +1944.6 2005.7 1951.9 2002.3 1959.2 1999 C +1960.1 1998.5 1960.1 1999.8 1960.4 2000.4 C +1959.7 2006.9 1959.7 2014.2 1959.4 2021.1 C +1959 2021.1 1959.2 2021.9 1959.2 2022.3 C +1959.2 2021.9 1959 2021.3 1959.4 2021.1 C +1959.8 2024.1 1959.2 2026.2 1962.5 2027.3 C +1963 2029.2 1965.3 2029.2 1965.9 2031.6 C +1968.3 2031.8 1967.8 2035.2 1971.2 2036 C +1970.8 2037.2 1971.9 2037.5 1972.8 2037.6 C +1974.9 2037.4 1973.9 2036.7 1974.3 2035.5 C +1973.3 2034.7 1974.1 2033.4 1973.6 2032.6 C +1973.9 2027.3 1973.9 2021.1 1973.6 2016.3 C +1973 2016 1973.9 2015.6 1974 2015.3 C +1974.3 2015.9 1975 2015.3 1975.2 2015.8 C +1975.3 2022.8 1975.1 2031.2 1975.5 2038.6 C +1977.9 2039 1973.7 2041.8 1976.2 2040 C +1975.7 2039 1975.5 2037.8 1975.5 2036.4 C +[0 0.4 1 0] vc +f +S +n +1991.1 2012.4 m +1987.5 2014.8 1983.4 2015.6 1979.8 2017.7 C +1978.5 2015.7 1981 2013.3 1978.1 2012.4 C +1973.6 2005.8 1966.8 2001.6 1962.3 1995.2 C +1961.4 1994.7 1960.8 1995 1960.1 1994.7 C +1962.5 1994.6 1964.6 1997.8 1966.1 1999.7 C +1969.7 2003.3 1974.2 2007.6 1977.6 2011.5 C +1977.5 2013.8 1980.6 2013.5 1979.1 2016.5 C +1979.1 2017 1979 2017.6 1979.3 2018 C +1983.1 2015.7 1987.8 2015.2 1991.1 2012.4 C +[0.18 0.18 0 0.78] vc +f +S +n +1970.9 2004.8 m +1970 2003.9 1969.2 2003 1968.3 2002.1 C +1969.2 2003 1970 2003.9 1970.9 2004.8 C +[0 0.87 0.91 0.83] vc +f +S +n +1887.9 1994.9 m +1888.5 1992.3 1891.4 1992.2 1893.2 1990.8 C +1898.4 1987.5 1904 1984.8 1909.5 1982.2 C +1909.7 1982.7 1910.3 1982.1 1910.4 1982.7 C +1909.5 1990.5 1910.1 1996.4 1910 2004.5 C +1909.1 2003.4 1909.7 2005.8 1909.5 2006.4 C +1910.4 2006 1909.7 2008 1910.2 2007.9 C +1911.3 2010.6 1912.5 2012.6 1915.7 2013.4 C +1915.8 2013.7 1915.5 2014.4 1916 2014.4 C +1916.3 2015 1915.4 2016 1915.2 2016 C +1916.1 2015.5 1916.5 2014.5 1916 2013.6 C +1913.4 2013.3 1913.1 2010.5 1910.9 2009.8 C +1910.7 2008.8 1910.4 2007.9 1910.2 2006.9 C +1911.1 1998.8 1909.4 1990.7 1910.7 1982.4 C +1910 1982.1 1908.9 1982.1 1908.3 1982.4 C +1901.9 1986.1 1895 1988.7 1888.8 1993 C +1888 1993.4 1888.4 1994.3 1887.6 1994.7 C +1888.1 2001.3 1887.8 2008.6 1887.9 2015.1 C +1887.3 2017.5 1887.9 2015.4 1888.4 2014.4 C +1887.8 2008 1888.4 2001.3 1887.9 1994.9 C +[0.07 0.06 0 0.58] vc +f +S +n +vmrs +1887.9 2018.4 m +1887.5 2016.9 1888.5 2016 1888.8 2014.8 C +1890.1 2014.8 1891.1 2016.6 1892.4 2015.3 C +1892.4 2014.4 1893.8 2012.9 1894.4 2012.4 C +1895.9 2012.4 1896.6 2013.9 1897.7 2012.7 C +1898.4 2011.7 1898.6 2010.4 1899.6 2009.8 C +1901.7 2009.9 1902.9 2010.4 1904 2009.1 C +1904.3 2007.4 1904 2007.6 1904.9 2007.2 C +1906.2 2007 1907.6 2006.5 1908.8 2006.7 C +1910.6 2008.2 1909.8 2011.5 1912.6 2012 C +1912.4 2013 1913.8 2012.7 1914 2013.2 C +1911.5 2011.1 1909.1 2007.9 1909.2 2004.3 C +1909.5 2003.5 1909.9 2004.9 1909.7 2004.3 C +1909.9 1996.2 1909.3 1990.5 1910.2 1982.7 C +1909.5 1982.6 1909.5 1982.6 1908.8 1982.7 C +1903.1 1985.7 1897 1987.9 1891.7 1992 C +1890.5 1993 1888.2 1992.9 1888.1 1994.9 C +1888.7 2001.4 1888.1 2008.4 1888.6 2014.8 C +1888.3 2016 1887.2 2016.9 1887.6 2018.4 C +1892.3 2023.9 1897.6 2027.9 1902.3 2033.3 C +1898 2028.2 1892.1 2023.8 1887.9 2018.4 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1910.9 1995.2 m +1910.4 1999.8 1911 2003.3 1910.9 2008.1 C +1910.9 2003.8 1910.9 1999.2 1910.9 1995.2 C +[0.18 0.18 0 0.78] vc +f +S +n +1911.2 2004.3 m +1911.2 2001.9 1911.2 1999.7 1911.2 1997.3 C +1911.2 1999.7 1911.2 2001.9 1911.2 2004.3 C +[0 0.87 0.91 0.83] vc +f +S +n +1958.7 1995.2 m +1959 1995.6 1956.2 1995 1956.5 1996.8 C +1955.8 1997.6 1954.2 1998.5 1953.6 1997.3 C +1953.6 1990.8 1954.9 1989.6 1953.4 1983.9 C +1953.4 1983.3 1953.3 1982.1 1954.4 1982 C +1955.5 1982.6 1956.5 1981.3 1957.5 1981 C +1956.3 1981.8 1954.7 1982.6 1953.9 1981.5 C +1951.4 1983 1954.7 1988.8 1952.9 1990.6 C +1953.8 1990.6 1953.2 1992.7 1953.4 1993.7 C +1953.8 1994.5 1952.3 1996.1 1953.2 1997.8 C +1956.3 1999.4 1957.5 1994 1959.9 1995.6 C +1962 1994.4 1963.7 1997.7 1965.2 1998.8 C +1963.5 1996.7 1961.2 1994.1 1958.7 1995.2 C +f +S +n +1945 2000.7 m +1945.4 1998.7 1945.4 1997.9 1945 1995.9 C +1944.5 1995.3 1944.2 1992.6 1945.7 1993.2 C +1946 1992.2 1948.7 1992.5 1948.4 1990.6 C +1947.5 1990.3 1948.1 1988.7 1947.9 1988.2 C +1948.9 1987.8 1950.5 1986.8 1950.5 1984.6 C +1951.5 1980.9 1946.7 1983 1947.2 1979.8 C +1944.5 1979.9 1945.2 1976.6 1943.1 1976.7 C +1941.8 1975.7 1942.1 1972.7 1939.2 1973.8 C +1938.2 1974.6 1939.3 1971.6 1938.3 1970.9 C +1938.8 1969.2 1933.4 1970.3 1937.3 1970 C +1939.4 1971.2 1937.2 1973 1937.6 1974.3 C +1937.2 1976.3 1937.1 1981.2 1937.8 1984.1 C +1938.8 1982.3 1937.9 1976.6 1938.5 1973.1 C +1938.9 1975 1938.5 1976.4 1939.7 1977.2 C +1939.5 1983.5 1938.9 1991.3 1940.2 1997.3 C +1939.4 1999.1 1938.6 1997.1 1937.8 1997.1 C +1937.4 1996.7 1937.6 1996.1 1937.6 1995.6 C +1936.5 1998.5 1940.1 1998.4 1940.9 2000.7 C +1942.1 2000.4 1943.2 2001.3 1943.1 2002.4 C +1943.6 2003.1 1941.1 2004.6 1942.8 2003.8 C +1943.9 2002.5 1942.6 2000.6 1945 2000.7 C +[0.65 0.65 0 0.42] vc +f +S +n +1914.5 2006.4 m +1914.1 2004.9 1915.2 2004 1915.5 2002.8 C +1916.7 2002.8 1917.8 2004.6 1919.1 2003.3 C +1919 2002.4 1920.4 2000.9 1921 2000.4 C +1922.5 2000.4 1923.2 2001.9 1924.4 2000.7 C +1925 1999.7 1925.3 1998.4 1926.3 1997.8 C +1928.4 1997.9 1929.5 1998.4 1930.6 1997.1 C +1930.9 1995.4 1930.7 1995.6 1931.6 1995.2 C +1932.8 1995 1934.3 1994.5 1935.4 1994.7 C +1936.1 1995.8 1936.9 1996.2 1936.6 1997.8 C +1938.9 1999.4 1939.7 2001.3 1942.4 2002.4 C +1942.4 2002.5 1942.2 2003 1942.6 2002.8 C +1942.9 2000.4 1939.2 2001.8 1939.2 1999.7 C +1936.2 1998.6 1937 1995.3 1935.9 1993.5 C +1937.1 1986.5 1935.2 1977.9 1937.6 1971.2 C +1937.6 1970.3 1936.6 1971 1936.4 1970.4 C +1930.2 1973.4 1924 1976 1918.4 1980 C +1917.2 1981 1914.9 1980.9 1914.8 1982.9 C +1915.3 1989.4 1914.7 1996.4 1915.2 2002.8 C +1914.9 2004 1913.9 2004.9 1914.3 2006.4 C +1919 2011.9 1924.2 2015.9 1928.9 2021.3 C +1924.6 2016.2 1918.7 2011.8 1914.5 2006.4 C +[0.4 0.4 0 0] vc +f +S +n +1914.5 1982.9 m +1915.1 1980.3 1918 1980.2 1919.8 1978.8 C +1925 1975.5 1930.6 1972.8 1936.1 1970.2 C +1939.4 1970.6 1936.1 1974.2 1936.6 1976.4 C +1936.5 1981.9 1936.8 1987.5 1936.4 1992.8 C +1935.9 1992.8 1936.2 1993.5 1936.1 1994 C +1937.1 1993.6 1936.2 1995.9 1936.8 1995.9 C +1937 1998 1939.5 1999.7 1940.4 2000.7 C +1940.1 1998.6 1935 1997.2 1937.6 1993.7 C +1938.3 1985.7 1935.9 1976.8 1937.8 1970.7 C +1936.9 1969.8 1935.4 1970.3 1934.4 1970.7 C +1928.3 1974.4 1921.4 1976.7 1915.5 1981 C +1914.6 1981.4 1915.1 1982.3 1914.3 1982.7 C +1914.7 1989.3 1914.5 1996.6 1914.5 2003.1 C +1913.9 2005.5 1914.5 2003.4 1915 2002.4 C +1914.5 1996 1915.1 1989.3 1914.5 1982.9 C +[0.07 0.06 0 0.58] vc +f +S +n +1939.2 1994.9 m +1939.3 1995 1939.4 1995.1 1939.5 1995.2 C +1939.1 1989 1939.3 1981.6 1939 1976.7 C +1938.6 1976.3 1938.6 1974.6 1938.5 1973.3 C +1938.7 1976.1 1938.1 1979.4 1939 1981.7 C +1937.3 1986 1937.7 1991.6 1938 1996.4 C +1937.3 1994.3 1939.6 1996.2 1939.2 1994.9 C +[0.18 0.18 0 0.78] vc +f +S +n +1938.3 1988.4 m +1938.5 1990.5 1937.9 1994.1 1938.8 1994.7 C +1937.9 1992.6 1939 1990.6 1938.3 1988.4 C +[0 0.87 0.91 0.83] vc +f +S +n +1938.8 1985.8 m +1938.5 1985.9 1938.4 1985.7 1938.3 1985.6 C +1938.4 1986.2 1938 1989.5 1938.8 1987.2 C +1938.8 1986.8 1938.8 1986.3 1938.8 1985.8 C +f +S +n +vmrs +1972.8 2062.1 m +1971.9 2061 1972.5 2059.4 1972.4 2058 C +1972.2 2063.8 1971.9 2073.7 1972.4 2081.3 C +1972.5 2074.9 1971.9 2067.9 1972.8 2062.1 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +S +n +1940.2 2071.7 m +1941.3 2072 1943.1 2072.3 1944 2071.5 C +1943.6 2069.9 1945.2 2069.1 1946 2068.8 C +1950 2071.1 1948.7 2065.9 1951.7 2066.2 C +1953.5 2063.9 1956.9 2069.4 1955.6 2063.8 C +1955.5 2064.2 1955.7 2064.8 1955.3 2065 C +1954.3 2063.7 1956.2 2063.6 1955.6 2062.1 C +1954.5 2060 1958.3 2050.3 1952.2 2055.6 C +1949.1 2053.8 1946 2051 1943.8 2048 C +1940.3 2048 1937.5 2051.3 1934.2 2052.5 C +1933.1 2054.6 1934.4 2057.3 1934 2060 C +1934 2065.1 1934 2069.7 1934 2074.6 C +1934.4 2069 1934.1 2061.5 1934.2 2054.9 C +1934.6 2054.5 1935.3 2054.7 1935.9 2054.7 C +1937 2055.3 1935.9 2056.1 1935.9 2056.8 C +1936.5 2063 1935.6 2070.5 1935.9 2074.6 C +1936.7 2074.4 1937.3 2075.2 1938 2074.6 C +1937.9 2073.6 1939.1 2072.1 1940.2 2071.7 C +[0 0.2 1 0] vc +f +S +n +1933.2 2074.1 m +1933.2 2071.5 1933.2 2069 1933.2 2066.4 C +1933.2 2069 1933.2 2071.5 1933.2 2074.1 C +[0 1 1 0.36] vc +f +S +n +2007.4 2048.9 m +2006.5 2047.8 2007.1 2046.2 2006.9 2044.8 C +2006.7 2050.6 2006.5 2060.5 2006.9 2068.1 C +2007.1 2061.7 2006.5 2054.7 2007.4 2048.9 C +f +S +n +1927.2 2062.4 m +1925.8 2060.1 1928.1 2058.2 1927 2056.4 C +1927.3 2055.5 1926.5 2053.5 1926.8 2051.8 C +1926.8 2052.8 1926 2052.5 1925.3 2052.5 C +1924.1 2052.8 1925 2050.5 1924.4 2050.1 C +1925.3 2050.2 1925.4 2048.8 1926.3 2049.4 C +1926.5 2052.3 1928.4 2047.2 1928.4 2051.1 C +1928.9 2050.5 1929 2051.4 1928.9 2051.8 C +1928.9 2052 1928.9 2052.3 1928.9 2052.5 C +1929.4 2051.4 1928.9 2049 1930.1 2048.2 C +1928.9 2047.1 1930.5 2047.1 1930.4 2046.5 C +1931.9 2046.2 1933.1 2046.1 1934.7 2046.5 C +1934.6 2046.9 1935.2 2047.9 1934.4 2048.4 C +1936.9 2048.1 1933.6 2043.8 1935.9 2043.9 C +1935.7 2043.9 1934.8 2041.3 1933.2 2041.7 C +1932.5 2041.6 1932.4 2039.6 1932.3 2041 C +1930.8 2042.6 1929 2040.6 1927.7 2042 C +1927.5 2041.4 1927.1 2040.9 1927.2 2040.3 C +1927.8 2040.6 1927.4 2039.1 1928.2 2038.6 C +1929.4 2038 1930.5 2038.8 1931.3 2037.9 C +1931.7 2039 1932.5 2038.6 1931.8 2037.6 C +1930.9 2037 1928.7 2037.8 1928.2 2037.9 C +1926.7 2037.8 1928 2039 1927 2038.8 C +1927.4 2040.4 1925.6 2040.8 1925.1 2041 C +1924.3 2040.4 1923.2 2040.5 1922.2 2040.5 C +1921.4 2041.7 1921 2043.9 1919.3 2043.9 C +1918.8 2043.4 1917.2 2043.3 1916.4 2043.4 C +1915.9 2044.4 1915.7 2046 1914.3 2046.5 C +1913.1 2046.6 1912 2044.5 1911.4 2046.3 C +1912.8 2046.5 1913.8 2047.4 1915.7 2047 C +1916.9 2047.7 1915.6 2048.8 1916 2049.4 C +1915.4 2049.3 1913.9 2050.3 1913.3 2051.1 C +1913.9 2054.1 1916 2050.2 1916.7 2053 C +1916.9 2053.8 1915.5 2054.1 1916.7 2054.4 C +1917 2054.7 1920.2 2054.3 1919.3 2056.6 C +1918.8 2056.1 1920.2 2058.6 1920.3 2057.6 C +1921.2 2057.9 1922.1 2057.5 1922.4 2059 C +1922.3 2059.1 1922.2 2059.3 1922 2059.2 C +1922.1 2059.7 1922.4 2060.3 1922.9 2060.7 C +1923.2 2060.1 1923.8 2060.4 1924.6 2060.7 C +1925.9 2062.6 1923.2 2062 1925.6 2063.6 C +1926.1 2063.1 1927.3 2062.5 1927.2 2062.4 C +[0.21 0.21 0 0] vc +f +S +n +1933.2 2063.3 m +1933.2 2060.7 1933.2 2058.2 1933.2 2055.6 C +1933.2 2058.2 1933.2 2060.7 1933.2 2063.3 C +[0 1 1 0.36] vc +f +S +n +1965.2 2049.2 m +1967.1 2050.1 1969.9 2053.7 1972.1 2056.4 C +1970.5 2054 1967.6 2051.3 1965.2 2049.2 C +f +S +n +1991.8 2034.8 m +1991.7 2041.5 1992 2048.5 1991.6 2055.2 C +1990.5 2056.4 1991.9 2054.9 1991.8 2054.4 C +1991.8 2047.9 1991.8 2041.3 1991.8 2034.8 C +f +S +n +1988.9 2053.2 m +1988.9 2044.3 1988.9 2036.6 1988.9 2028.3 C +1985.7 2028.2 1987.2 2023.5 1983.9 2024.2 C +1983.9 2022.4 1982 2021.6 1981 2021.3 C +1980.6 2021.1 1980.6 2021.7 1980.3 2021.6 C +1980.3 2027 1980.3 2034.8 1980.3 2041.5 C +1979.3 2043.2 1977.6 2043 1976.2 2043.6 C +1977.1 2043.8 1978.5 2043.2 1978.8 2044.1 C +1978.5 2045.3 1979.9 2045.3 1980.3 2045.8 C +1980.5 2046.8 1980.7 2046.2 1981.5 2046.5 C +1982.4 2047.1 1982 2048.6 1982.7 2049.4 C +1984.2 2049.6 1984.6 2052.2 1986.8 2051.6 C +1987.1 2048.6 1985.1 2042.7 1986.5 2040.5 C +1986.3 2036.7 1986.9 2031.7 1986 2029.2 C +1986.3 2027.1 1986.9 2028.6 1987.7 2027.6 C +1987.7 2028.3 1988.7 2028 1988.7 2028.8 C +1988.1 2033 1988.7 2037.5 1988.2 2041.7 C +1987.8 2041.4 1988 2040.8 1988 2040.3 C +1988 2041 1988 2042.4 1988 2042.4 C +1988 2042.4 1988.1 2042.3 1988.2 2042.2 C +1989.3 2046 1988 2050.2 1988.4 2054 C +1987.8 2054.4 1987.1 2054.7 1986.5 2055.4 C +1987.4 2054.4 1988.4 2054.6 1988.9 2053.2 C +[0 1 1 0.23] vc +f +S +n +1950.8 2054.4 m +1949.7 2053.4 1948.7 2052.3 1947.6 2051.3 C +1948.7 2052.3 1949.7 2053.4 1950.8 2054.4 C +[0 1 1 0.36] vc +f +S +n +vmrs +2006.7 2043.2 m +2004.5 2040.8 2002.4 2038.4 2000.2 2036 C +2002.4 2038.4 2004.5 2040.8 2006.7 2043.2 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +S +n +1976.7 2019.6 m +1975.8 2018.6 1976.4 2016.9 1976.2 2015.6 C +1976 2021.3 1975.8 2031.2 1976.2 2038.8 C +1976.4 2032.4 1975.8 2025.5 1976.7 2019.6 C +f +S +n +1988.4 2053.5 m +1988.6 2049.2 1988.1 2042.8 1988 2040 C +1988.4 2040.4 1988.1 2041 1988.2 2041.5 C +1988.3 2037.2 1988 2032.7 1988.4 2028.5 C +1987.6 2027.1 1987.2 2028.6 1986.8 2028 C +1985.9 2028.5 1986.5 2029.7 1986.3 2030.4 C +1986.9 2029.8 1986.6 2031 1987 2031.2 C +1987.4 2039.6 1985 2043 1987.2 2050.4 C +1987.2 2051.6 1985.9 2052.3 1984.6 2051.3 C +1981.9 2049.7 1982.9 2047 1980.3 2046.5 C +1980.3 2045.2 1978.1 2046.2 1978.6 2043.9 C +1975.6 2043.3 1979.3 2045.6 1979.6 2046.5 C +1980.8 2046.6 1981.5 2048.5 1982.2 2049.9 C +1983.7 2050.8 1984.8 2052.8 1986.5 2053 C +1986.7 2053.5 1987.5 2054.1 1987 2054.7 C +1987.4 2053.9 1988.3 2054.3 1988.4 2053.5 C +[0 1 1 0.23] vc +f +S +n +1988 2038.1 m +1988 2036.7 1988 2035.4 1988 2034 C +1988 2035.4 1988 2036.7 1988 2038.1 C +[0 1 1 0.36] vc +f +S +n +1999.7 2035.7 m +1997.6 2033.5 1995.4 2031.2 1993.2 2029 C +1995.4 2031.2 1997.6 2033.5 1999.7 2035.7 C +f +S +n +1944 2029.2 m +1945.2 2029.5 1946.9 2029.8 1947.9 2029 C +1947.4 2027.4 1949 2026.7 1949.8 2026.4 C +1953.9 2028.6 1952.6 2023.4 1955.6 2023.7 C +1957.4 2021.4 1960.7 2027 1959.4 2021.3 C +1959.3 2021.7 1959.6 2022.3 1959.2 2022.5 C +1958.1 2021.2 1960.1 2021.1 1959.4 2019.6 C +1959.1 2012.7 1959.9 2005.1 1959.6 1999.2 C +1955.3 2000.1 1951.3 2003.1 1947.2 2005 C +1943.9 2006 1941.2 2008.7 1938 2010 C +1936.9 2012.1 1938.2 2014.8 1937.8 2017.5 C +1937.8 2022.6 1937.8 2027.3 1937.8 2032.1 C +1938.2 2026.5 1938 2019 1938 2012.4 C +1938.5 2012 1939.2 2012.3 1939.7 2012.2 C +1940.8 2012.8 1939.7 2013.6 1939.7 2014.4 C +1940.4 2020.5 1939.4 2028 1939.7 2032.1 C +1940.6 2031.9 1941.2 2032.7 1941.9 2032.1 C +1941.7 2031.2 1943 2029.7 1944 2029.2 C +[0 0.2 1 0] vc +f +S +n +1937.1 2031.6 m +1937.1 2029.1 1937.1 2026.5 1937.1 2024 C +1937.1 2026.5 1937.1 2029.1 1937.1 2031.6 C +[0 1 1 0.36] vc +f +S +n +1991.8 2028 m +1992.5 2027.8 1993.2 2029.9 1994 2030.2 C +1992.9 2029.6 1993.1 2028.1 1991.8 2028 C +[0 1 1 0.23] vc +f +S +n +1991.8 2027.8 m +1992.4 2027.6 1992.6 2028.3 1993 2028.5 C +1992.6 2028.2 1992.2 2027.6 1991.6 2027.8 C +1991.6 2028.5 1991.6 2029.1 1991.6 2029.7 C +1991.6 2029.1 1991.4 2028.3 1991.8 2027.8 C +[0 1 1 0.36] vc +f +S +n +1985.8 2025.4 m +1985.3 2025.2 1984.8 2024.7 1984.1 2024.9 C +1983.3 2025.3 1983.6 2027.3 1983.9 2027.6 C +1985 2028 1986.9 2026.9 1985.8 2025.4 C +[0 1 1 0.23] vc +f +S +n +vmrs +1993.5 2024.4 m +1992.4 2023.7 1991.3 2022.9 1990.1 2023.2 C +1990.7 2023.7 1989.8 2023.8 1989.4 2023.7 C +1989.1 2023.7 1988.6 2023.9 1988.4 2023.5 C +1988.5 2023.2 1988.3 2022.7 1988.7 2022.5 C +1989 2022.6 1988.9 2023 1988.9 2023.2 C +1989.1 2022.8 1990.4 2022.3 1990.6 2021.3 C +1990.4 2021.8 1990 2021.3 1990.1 2021.1 C +1990.1 2020.9 1990.1 2020.1 1990.1 2020.6 C +1989.9 2021.1 1989.5 2020.6 1989.6 2020.4 C +1989.6 2019.8 1988.7 2019.6 1988.2 2019.2 C +1987.5 2018.7 1987.7 2020.2 1987 2019.4 C +1987.5 2020.4 1986 2021.1 1987.5 2021.8 C +1986.8 2023.1 1986.6 2021.1 1986 2021.1 C +1986.1 2020.1 1985.9 2019 1986.3 2018.2 C +1986.7 2018.4 1986.5 2019 1986.5 2019.4 C +1986.5 2018.7 1986.4 2017.8 1987.2 2017.7 C +1986.5 2017.2 1985.5 2019.3 1985.3 2020.4 C +1986.2 2022 1987.3 2023.5 1989.2 2024.2 C +1990.8 2024.3 1991.6 2022.9 1993.2 2024.4 C +1993.8 2025.4 1995 2026.6 1995.9 2027.1 C +1995 2026.5 1994.1 2025.5 1993.5 2024.4 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +[0 0.5 0.5 0.2] vc +S +n +2023 2040.3 m +2023.2 2036 2022.7 2029.6 2022.5 2026.8 C +2022.9 2027.2 2022.7 2027.8 2022.8 2028.3 C +2022.8 2024 2022.6 2019.5 2023 2015.3 C +2022.2 2013.9 2021.7 2015.4 2021.3 2014.8 C +2020.4 2015.3 2021 2016.5 2020.8 2017.2 C +2021.4 2016.6 2021.1 2017.8 2021.6 2018 C +2022 2026.4 2019.6 2029.8 2021.8 2037.2 C +2021.7 2038.4 2020.5 2039.1 2019.2 2038.1 C +2016.5 2036.5 2017.5 2033.8 2014.8 2033.3 C +2014.9 2032 2012.6 2033 2013.2 2030.7 C +2011.9 2030.8 2011.2 2030.1 2010.8 2029.2 C +2010.8 2029.1 2010.8 2028.2 2010.8 2028.8 C +2010 2028.8 2010.4 2026.5 2008.6 2027.3 C +2007.9 2026.6 2007.3 2025.9 2007.9 2027.1 C +2009.7 2028 2010 2030.1 2012.2 2030.9 C +2012.9 2032.1 2013.7 2033.6 2015.1 2033.6 C +2015.7 2035.1 2016.9 2036.7 2018.4 2038.4 C +2019.8 2039.3 2022 2039.4 2021.6 2041.5 C +2021.9 2040.7 2022.9 2041.1 2023 2040.3 C +[0 1 1 0.23] vc +f +S +n +2022.5 2024.9 m +2022.5 2023.5 2022.5 2022.2 2022.5 2020.8 C +2022.5 2022.2 2022.5 2023.5 2022.5 2024.9 C +[0 1 1 0.36] vc +f +S +n +1983.2 2022.8 m +1982.4 2022.5 1982.1 2021.6 1981.2 2022.3 C +1981.1 2022.9 1980.5 2024 1981 2024.2 C +1981.8 2024.6 1982.9 2024.4 1983.2 2022.8 C +[0 1 1 0.23] vc +f +S +n +1931.1 2019.9 m +1929.6 2017.7 1932 2015.7 1930.8 2013.9 C +1931.1 2013 1930.3 2011 1930.6 2009.3 C +1930.6 2010.3 1929.8 2010 1929.2 2010 C +1928 2010.3 1928.8 2008.1 1928.2 2007.6 C +1929.1 2007.8 1929.3 2006.3 1930.1 2006.9 C +1930.3 2009.8 1932.2 2004.8 1932.3 2008.6 C +1932.7 2008 1932.8 2009 1932.8 2009.3 C +1932.8 2009.6 1932.8 2009.8 1932.8 2010 C +1933.2 2009 1932.7 2006.6 1934 2005.7 C +1932.7 2004.6 1934.3 2004.6 1934.2 2004 C +1935.8 2003.7 1937 2003.6 1938.5 2004 C +1938.5 2004.5 1939.1 2005.4 1938.3 2006 C +1940.7 2005.7 1937.4 2001.3 1939.7 2001.4 C +1939.5 2001.4 1938.6 1998.8 1937.1 1999.2 C +1936.3 1999.1 1936.2 1997.1 1936.1 1998.5 C +1934.7 2000.1 1932.9 1998.2 1931.6 1999.5 C +1931.3 1998.9 1930.9 1998.5 1931.1 1997.8 C +1931.6 1998.2 1931.3 1996.6 1932 1996.1 C +1933.2 1995.5 1934.3 1996.4 1935.2 1995.4 C +1935.5 1996.5 1936.3 1996.1 1935.6 1995.2 C +1934.7 1994.5 1932.5 1995.3 1932 1995.4 C +1930.5 1995.3 1931.9 1996.5 1930.8 1996.4 C +1931.2 1997.9 1929.5 1998.3 1928.9 1998.5 C +1928.1 1997.9 1927.1 1998 1926 1998 C +1925.3 1999.2 1924.8 2001.4 1923.2 2001.4 C +1922.6 2000.9 1921 2000.9 1920.3 2000.9 C +1919.7 2001.9 1919.6 2003.5 1918.1 2004 C +1916.9 2004.1 1915.8 2002 1915.2 2003.8 C +1916.7 2004 1917.6 2004.9 1919.6 2004.5 C +1920.7 2005.2 1919.4 2006.3 1919.8 2006.9 C +1919.2 2006.9 1917.7 2007.8 1917.2 2008.6 C +1917.8 2011.6 1919.8 2007.8 1920.5 2010.5 C +1920.8 2011.3 1919.3 2011.6 1920.5 2012 C +1920.8 2012.3 1924 2011.8 1923.2 2014.1 C +1922.6 2013.6 1924.1 2016.1 1924.1 2015.1 C +1925.1 2015.4 1925.9 2015 1926.3 2016.5 C +1926.2 2016.6 1926 2016.8 1925.8 2016.8 C +1925.9 2017.2 1926.2 2017.8 1926.8 2018.2 C +1927.1 2017.6 1927.7 2018 1928.4 2018.2 C +1929.7 2020.1 1927.1 2019.5 1929.4 2021.1 C +1929.9 2020.7 1931.1 2020 1931.1 2019.9 C +[0.21 0.21 0 0] vc +f +S +n +1937.1 2020.8 m +1937.1 2018.3 1937.1 2015.7 1937.1 2013.2 C +1937.1 2015.7 1937.1 2018.3 1937.1 2020.8 C +[0 1 1 0.36] vc +f +S +n +2020.4 2012.2 m +2019.8 2012 2019.3 2011.5 2018.7 2011.7 C +2017.9 2012.1 2018.1 2014.1 2018.4 2014.4 C +2019.6 2014.8 2021.4 2013.7 2020.4 2012.2 C +[0 1 1 0.23] vc +f +S +n +1976 2013.9 m +1973.8 2011.5 1971.6 2009.1 1969.5 2006.7 C +1971.6 2009.1 1973.8 2011.5 1976 2013.9 C +[0 1 1 0.36] vc +f +S +n +1995.4 2012.7 m +1996.1 2010.3 1993.8 2006.2 1997.3 2005.7 C +1998.9 2005.4 2000 2003.7 2001.4 2003.1 C +2003.9 2003.1 2005.3 2001.3 2006.9 1999.7 C +2004.5 2003.5 2000 2002.2 1997.6 2005.7 C +1996.5 2005.9 1994.8 2006.1 1995.2 2007.6 C +1995.7 2009.4 1995.2 2011.6 1994.7 2012.9 C +1992 2015.8 1987.8 2015.7 1985.3 2018.7 C +1988.3 2016.3 1992.3 2015.3 1995.4 2012.7 C +[0.18 0.18 0 0.78] vc +f +S +n +1995.6 2012.4 m +1995.6 2011.2 1995.6 2010 1995.6 2008.8 C +1995.6 2010 1995.6 2011.2 1995.6 2012.4 C +[0 1 1 0.36] vc +f +S +n +vmrs +2017.7 2009.6 m +2016.9 2009.3 2016.7 2008.4 2015.8 2009.1 C +2014.2 2010.6 2016 2010.6 2016.5 2011.5 C +2017.2 2010.9 2018.1 2010.8 2017.7 2009.6 C +[0 1 1 0.23] vc +f +0.4 w +2 J +2 M +S +n +2014.4 2006.4 m +2013.5 2006.8 2012.1 2005.6 2012 2006.7 C +2013 2007.3 2011.9 2009.2 2012.9 2008.4 C +2014.2 2008.3 2014.6 2007.8 2014.4 2006.4 C +f +S +n +1969 2006.4 m +1966.5 2003.8 1964 2001.2 1961.6 1998.5 C +1964 2001.2 1966.5 2003.8 1969 2006.4 C +[0 1 1 0.36] vc +f +S +n +2012 2005.2 m +2012.2 2004.2 2011.4 2003.3 2010.3 2003.3 C +2009 2003.6 2010 2004.7 2009.6 2004.8 C +2009.3 2005.7 2011.4 2006.7 2012 2005.2 C +[0 1 1 0.23] vc +f +S +n +1962.8 1995.2 m +1961.7 1994.4 1960.6 1993.7 1959.4 1994 C +1959.5 1994.9 1957.5 1994.1 1956.8 1994.7 C +1955.9 1995.5 1956.7 1997 1955.1 1997.3 C +1956.9 1996.7 1956.8 1994 1959.2 1994.7 C +1961.1 1991 1968.9 2003.2 1962.8 1995.2 C +[0 1 1 0.36] vc +f +S +n +1954.6 1995.6 m +1955.9 1994.7 1955.1 1989.8 1955.3 1988 C +1954.5 1988.3 1954.9 1986.6 1954.4 1986 C +1955.7 1989.2 1953.9 1991.1 1954.8 1994.2 C +1954.5 1995.9 1953.5 1995.3 1953.9 1997.3 C +1955.3 1998.3 1953.2 1995.5 1954.6 1995.6 C +f +S +n +1992.3 2011 m +1992.5 2006.7 1992 2000.3 1991.8 1997.6 C +1992.2 1997.9 1992 1998.5 1992 1999 C +1992.1 1994.7 1991.9 1990.2 1992.3 1986 C +1991.4 1984.6 1991 1986.1 1990.6 1985.6 C +1989.7 1986 1990.3 1987.2 1990.1 1988 C +1990.7 1987.4 1990.4 1988.5 1990.8 1988.7 C +1991.3 1997.1 1988.9 2000.6 1991.1 2007.9 C +1991 2009.1 1989.8 2009.9 1988.4 2008.8 C +1985.7 2007.2 1986.8 2004.5 1984.1 2004 C +1984.2 2002.7 1981.9 2003.7 1982.4 2001.4 C +1981.2 2001.5 1980.5 2000.8 1980 2000 C +1980 1999.8 1980 1998.9 1980 1999.5 C +1979.3 1999.5 1979.7 1997.2 1977.9 1998 C +1977.2 1997.3 1976.6 1996.7 1977.2 1997.8 C +1979 1998.7 1979.3 2000.8 1981.5 2001.6 C +1982.2 2002.8 1983 2004.3 1984.4 2004.3 C +1985 2005.8 1986.2 2007.5 1987.7 2009.1 C +1989 2010 1991.3 2010.2 1990.8 2012.2 C +1991.2 2011.4 1992.2 2011.8 1992.3 2011 C +[0 1 1 0.23] vc +f +S +n +1991.8 1995.6 m +1991.8 1994.3 1991.8 1992.9 1991.8 1991.6 C +1991.8 1992.9 1991.8 1994.3 1991.8 1995.6 C +[0 1 1 0.36] vc +f +S +n +1959.2 1994.2 m +1958.8 1993.3 1960.7 1993.9 1961.1 1993.7 C +1961.5 1993.9 1961.2 1994.4 1961.8 1994.2 C +1960.9 1994 1960.8 1992.9 1959.9 1992.5 C +1959.6 1993.5 1958.3 1993.5 1958.2 1994.2 C +1958.1 1994.1 1958 1994 1958 1994 C +1957.2 1994.9 1958 1993.4 1956.8 1993 C +1955.6 1992.5 1956 1991 1956.3 1989.9 C +1956.5 1989.8 1956.6 1990 1956.8 1990.1 C +1957.1 1989 1956 1989.1 1955.8 1988.2 C +1955.1 1990.4 1956.2 1995 1954.8 1995.9 C +1954.1 1995.5 1954.5 1996.5 1954.4 1997.1 C +1955 1996.8 1954.8 1997.4 1955.6 1996.8 C +1956 1996 1956.3 1993.2 1958.7 1994.2 C +1958.9 1994.2 1959.7 1994.2 1959.2 1994.2 C +[0 1 1 0.23] vc +f +S +n +1958.2 1994 m +1958.4 1993.5 1959.7 1993.1 1959.9 1992 C +1959.7 1992.5 1959.3 1992 1959.4 1991.8 C +1959.4 1991.6 1959.4 1990.8 1959.4 1991.3 C +1959.2 1991.8 1958.8 1991.3 1958.9 1991.1 C +1958.9 1990.5 1958 1990.3 1957.5 1989.9 C +1956.8 1989.5 1956.9 1991 1956.3 1990.1 C +1956.7 1991 1955.4 1992.1 1956.5 1992.3 C +1956.8 1993.5 1958.3 1992.9 1957.2 1994 C +1957.8 1994.3 1958.1 1992.4 1958.2 1994 C +[0 0.5 0.5 0.2] vc +f +S +n +vmrs +1954.4 1982.7 m +1956.1 1982.7 1954.1 1982.5 1953.9 1982.9 C +1953.9 1983.7 1953.7 1984.7 1954.1 1985.3 C +1954.4 1984.2 1953.6 1983.6 1954.4 1982.7 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +S +n +1989.6 1982.9 m +1989.1 1982.7 1988.6 1982.3 1988 1982.4 C +1987.2 1982.8 1987.4 1984.8 1987.7 1985.1 C +1988.9 1985.6 1990.7 1984.4 1989.6 1982.9 C +[0 1 1 0.23] vc +f +S +n +1987 1980.3 m +1986.2 1980 1986 1979.1 1985.1 1979.8 C +1983.5 1981.4 1985.3 1981.4 1985.8 1982.2 C +1986.5 1981.7 1987.4 1981.5 1987 1980.3 C +f +S +n +1983.6 1977.2 m +1982.7 1977.5 1981.4 1976.3 1981.2 1977.4 C +1982.3 1978 1981.2 1979.9 1982.2 1979.1 C +1983.5 1979 1983.9 1978.5 1983.6 1977.2 C +f +S +n +1981.2 1976 m +1981.5 1974.9 1980.6 1974 1979.6 1974 C +1978.3 1974.3 1979.3 1975.4 1978.8 1975.5 C +1978.6 1976.4 1980.7 1977.4 1981.2 1976 C +f +S +n +1972.1 2082.3 m +1971.8 2081.8 1971.3 2080.9 1971.2 2080.1 C +1971.1 2072.9 1971.3 2064.6 1970.9 2058.3 C +1970.3 2058.5 1970.1 2057.7 1969.7 2058.5 C +1970.6 2058.5 1969.7 2059 1970.2 2059.2 C +1970.2 2065.4 1970.2 2072.4 1970.2 2077.7 C +1971.1 2078.9 1970.6 2078.9 1970.4 2079.9 C +1969.2 2080.2 1968.2 2080.4 1967.3 2079.6 C +1966.8 2077.8 1963.4 2076.3 1963.5 2075.1 C +1961.5 2075.5 1962 2071.5 1959.6 2072 C +1959.2 2070 1956.5 2069.3 1955.8 2067.6 C +1956 2068.4 1955.3 2069.7 1956.5 2069.8 C +1958.6 2068.9 1958.1 2073.5 1960.1 2072.4 C +1960.7 2075.9 1964.7 2074.6 1964.2 2078 C +1967.2 2078.6 1967.9 2081.6 1970.7 2080.6 C +1970.3 2081.1 1971.5 2081.2 1971.9 2082.3 C +1967.2 2084.3 1962.9 2087.1 1958.2 2089 C +1962.9 2087 1967.4 2084.4 1972.1 2082.3 C +[0 0.2 1 0] vc +f +S +n +1971.9 2080.1 m +1971.9 2075.1 1971.9 2070 1971.9 2065 C +1971.9 2070 1971.9 2075.1 1971.9 2080.1 C +[0 1 1 0.23] vc +f +S +n +2010.8 2050.6 m +2013.2 2049 2010.5 2050.1 2010.5 2051.3 C +2010.5 2057.7 2010.5 2064.1 2010.5 2070.5 C +2008.7 2072.4 2006 2073.3 2003.6 2074.4 C +2016.4 2073.7 2008 2058.4 2010.8 2050.6 C +[0.4 0.4 0 0] vc +f +S +n +2006.4 2066.9 m +2006.4 2061.9 2006.4 2056.8 2006.4 2051.8 C +2006.4 2056.8 2006.4 2061.9 2006.4 2066.9 C +[0 1 1 0.23] vc +f +S +n +1971.9 2060.7 m +1972.2 2060.3 1971.4 2068.2 1972.4 2061.9 C +1971.8 2061.6 1972.4 2060.9 1971.9 2060.7 C +f +S +n +vmrs +1986.5 2055.2 m +1987.5 2054.3 1986.3 2053.4 1986 2052.8 C +1983.8 2052.7 1983.6 2050.1 1981.7 2049.6 C +1981.2 2048.7 1980.8 2047 1980.3 2046.8 C +1978.5 2047 1978 2044.6 1976.7 2043.9 C +1974 2044.4 1972 2046.6 1969.2 2047 C +1969 2047.2 1968.8 2047.5 1968.5 2047.7 C +1970.6 2049.6 1973.1 2051.3 1974.3 2054.2 C +1975.7 2054.5 1977 2055.2 1976.4 2057.1 C +1976.7 2058 1975.5 2058.5 1976 2059.5 C +1979.2 2058 1983 2056.6 1986.5 2055.2 C +[0 0.5 0.5 0.2] vc +f +0.4 w +2 J +2 M +S +n +1970.2 2054.2 m +1971.5 2055.3 1972.5 2056.8 1972.1 2058.3 C +1972.8 2056.5 1971.6 2055.6 1970.2 2054.2 C +[0 1 1 0.23] vc +f +S +n +1992 2052.5 m +1992 2053.4 1992.2 2054.4 1991.8 2055.2 C +1992.2 2054.4 1992 2053.4 1992 2052.5 C +f +S +n +1957.2 2053 m +1958.1 2052.6 1959 2052.2 1959.9 2051.8 C +1959 2052.2 1958.1 2052.6 1957.2 2053 C +f +S +n +2006.4 2047.5 m +2006.8 2047.1 2006 2055 2006.9 2048.7 C +2006.4 2048.4 2007 2047.7 2006.4 2047.5 C +f +S +n +2004.8 2041 m +2006.1 2042.1 2007.1 2043.6 2006.7 2045.1 C +2007.3 2043.3 2006.2 2042.4 2004.8 2041 C +f +S +n +1976 2039.8 m +1975.6 2039.3 1975.2 2038.4 1975 2037.6 C +1974.9 2030.4 1975.2 2022.1 1974.8 2015.8 C +1974.2 2016 1974 2015.3 1973.6 2016 C +1974.4 2016 1973.5 2016.5 1974 2016.8 C +1974 2022.9 1974 2030 1974 2035.2 C +1974.9 2036.4 1974.4 2036.4 1974.3 2037.4 C +1973.1 2037.7 1972 2037.9 1971.2 2037.2 C +1970.6 2035.3 1967.3 2033.9 1967.3 2032.6 C +1965.3 2033 1965.9 2029.1 1963.5 2029.5 C +1963 2027.6 1960.4 2026.8 1959.6 2025.2 C +1959.8 2025.9 1959.2 2027.2 1960.4 2027.3 C +1962.5 2026.4 1961.9 2031 1964 2030 C +1964.6 2033.4 1968.5 2032.1 1968 2035.5 C +1971 2036.1 1971.8 2039.1 1974.5 2038.1 C +1974.2 2038.7 1975.3 2038.7 1975.7 2039.8 C +1971 2041.8 1966.7 2044.6 1962 2046.5 C +1966.8 2044.5 1971.3 2041.9 1976 2039.8 C +[0 0.2 1 0] vc +f +S +n +1975.7 2037.6 m +1975.7 2032.6 1975.7 2027.6 1975.7 2022.5 C +1975.7 2027.6 1975.7 2032.6 1975.7 2037.6 C +[0 1 1 0.23] vc +f +S +n +1992 2035.5 m +1992 2034.2 1992 2032.9 1992 2031.6 C +1992 2032.9 1992 2034.2 1992 2035.5 C +f +S +n +2015.3 2036 m +2015.4 2034.1 2013.3 2034 2012.9 2033.3 C +2011.5 2031 2009.3 2029.4 2007.4 2028 C +2006.9 2027.1 2006.6 2023.8 2005 2024.9 C +2004 2024.9 2002.9 2024.9 2001.9 2024.9 C +2001.4 2026.5 2001 2028.4 2003.8 2028.3 C +2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C +2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C +2013 2035.5 2015.3 2037.4 2015.3 2036 C +[0 0 0 0] vc +f +S +n +vmrs +2009.1 2030.4 m +2009.1 2029 2007.5 2029.4 2006.9 2028.3 C +2007.2 2027.1 2006.5 2025.5 2005.7 2024.7 C +2004.6 2025.1 2003.1 2024.9 2001.9 2024.9 C +2001.8 2026.2 2000.9 2027 2002.4 2028 C +2004.5 2027.3 2004.9 2029.4 2006.9 2029 C +2007 2030.2 2007.6 2030.7 2008.4 2031.4 C +2008.8 2031.5 2009.1 2031.1 2009.1 2030.4 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +2003.8 2029.5 m +2003 2029.4 2001.9 2029.1 2002.4 2030.4 C +2003.1 2031.3 2005.2 2030.3 2003.8 2029.5 C +[0 1 1 0.23] vc +f +S +n +1999.2 2025.2 m +1999.1 2025.6 1998 2025.7 1998.8 2026.6 C +2000.9 2028.5 1999.5 2023.4 1999.2 2025.2 C +f +S +n +2007.6 2024.2 m +2007.6 2022.9 2008.4 2024.2 2007.6 2022.8 C +2007.6 2017.5 2007.8 2009.1 2007.4 2003.8 C +2007.9 2003.7 2008.7 2002.8 2009.1 2002.1 C +2009.6 2000.8 2008.3 2000.8 2007.9 2000.2 C +2004.9 2000 2008.9 2001.3 2007.2 2002.1 C +2006.7 2007.7 2007 2015.1 2006.9 2021.1 C +2006.7 2022.1 2005.4 2022.8 2006.2 2023.5 C +2006.6 2023.1 2008 2025.9 2007.6 2024.2 C +f +S +n +1989.9 2023.5 m +1989.5 2022.6 1991.4 2023.2 1991.8 2023 C +1992.2 2023.2 1991.9 2023.7 1992.5 2023.5 C +1991.6 2023.2 1991.6 2022.2 1990.6 2021.8 C +1990.4 2022.8 1989 2022.8 1988.9 2023.5 C +1988.5 2023 1988.7 2022.6 1988.7 2023.5 C +1989.1 2023.5 1990.2 2023.5 1989.9 2023.5 C +f +[0 0.5 0.5 0.2] vc +S +n +2003.3 2023.5 m +2003.1 2023.3 2003.1 2023.2 2003.3 2023 C +2003.7 2023.1 2003.9 2022.9 2003.8 2022.5 C +2003.4 2022.2 2001.2 2022.3 2002.4 2023 C +2002.6 2022.9 2002.7 2023.1 2002.8 2023.2 C +2000.7 2023.7 2003.9 2023.4 2003.3 2023.5 C +[0 1 1 0.23] vc +f +S +n +1986.8 2019.4 m +1987.8 2019.8 1987.5 2018.6 1987.2 2018 C +1986.2 2017.8 1987.3 2020.5 1986.3 2019.2 C +1986.3 2017.7 1986.3 2020.6 1986.3 2021.3 C +1988.5 2023.1 1985.6 2020.3 1986.8 2019.4 C +f +S +n +1975.7 2018.2 m +1976.1 2017.8 1975.2 2025.7 1976.2 2019.4 C +1975.7 2019.2 1976.3 2018.4 1975.7 2018.2 C +f +S +n +1974 2011.7 m +1975.4 2012.8 1976.4 2014.3 1976 2015.8 C +1976.6 2014 1975.5 2013.1 1974 2011.7 C +f +S +n +1984.6 2006.7 m +1984.7 2004.8 1982.6 2004.8 1982.2 2004 C +1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C +1976.1 1997.8 1975.8 1994.5 1974.3 1995.6 C +1973.3 1995.6 1972.2 1995.6 1971.2 1995.6 C +1970.7 1997.2 1970.3 1999.1 1973.1 1999 C +1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C +1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C +1982.3 2006.2 1984.5 2008.1 1984.6 2006.7 C +[0 0 0 0] vc +f +S +n +vmrs +1978.4 2001.2 m +1978.4 1999.7 1976.8 2000.1 1976.2 1999 C +1976.5 1997.8 1975.8 1996.2 1975 1995.4 C +1973.9 1995.8 1972.4 1995.6 1971.2 1995.6 C +1971 1997 1970.2 1997.7 1971.6 1998.8 C +1973.8 1998 1974.2 2000.1 1976.2 1999.7 C +1976.3 2000.9 1976.9 2001.4 1977.6 2002.1 C +1978.1 2002.2 1978.4 2001.8 1978.4 2001.2 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1973.1 2000.2 m +1972.3 2000.1 1971.2 1999.8 1971.6 2001.2 C +1972.4 2002 1974.5 2001 1973.1 2000.2 C +[0 1 1 0.23] vc +f +S +n +1960.8 1998.5 m +1961.6 1998.2 1962.6 2000.3 1963.2 2000.9 C +1962.3 2000.1 1962.2 1998.7 1960.8 1998.5 C +f +S +n +1968.5 1995.9 m +1968.4 1996.4 1967.3 1996.4 1968 1997.3 C +1970.1 1999.2 1968.8 1994.1 1968.5 1995.9 C +f +S +n +1976.9 1994.9 m +1976.9 1993.7 1977.6 1994.9 1976.9 1993.5 C +1976.9 1988.2 1977.1 1979.8 1976.7 1974.5 C +1977.2 1974.5 1978 1973.5 1978.4 1972.8 C +1978.8 1971.5 1977.6 1971.5 1977.2 1970.9 C +1974.2 1970.7 1978.2 1972 1976.4 1972.8 C +1976 1978.4 1976.3 1985.8 1976.2 1991.8 C +1976 1992.8 1974.6 1993.5 1975.5 1994.2 C +1975.9 1993.8 1977.3 1996.6 1976.9 1994.9 C +f +S +n +1972.6 1994.2 m +1972.4 1994 1972.4 1993.9 1972.6 1993.7 C +1973 1993.8 1973.1 1993.7 1973.1 1993.2 C +1972.7 1992.9 1970.5 1993.1 1971.6 1993.7 C +1971.9 1993.7 1972 1993.8 1972.1 1994 C +1970 1994.4 1973.1 1994.1 1972.6 1994.2 C +f +S +n +1948.1 2093.8 m +1947 2092.7 1945.9 2091.6 1944.8 2090.4 C +1945.9 2091.6 1947 2092.7 1948.1 2093.8 C +[0 0.4 1 0] vc +f +S +n +1953.4 2091.4 m +1954.8 2090.7 1956.3 2090 1957.7 2089.2 C +1956.3 2090 1954.8 2090.7 1953.4 2091.4 C +[0 0.2 1 0] vc +f +S +n +1954.1 2091.4 m +1956.6 2089.6 1957.2 2089.6 1954.1 2091.4 C +[0 0.4 1 0] vc +f +S +n +1962.3 2087.3 m +1963.7 2086.6 1965.2 2085.9 1966.6 2085.2 C +1965.2 2085.9 1963.7 2086.6 1962.3 2087.3 C +f +S +n +vmrs +1967.1 2084.9 m +1968.3 2084.4 1969.7 2083.8 1970.9 2083.2 C +1969.7 2083.8 1968.3 2084.4 1967.1 2084.9 C +[0 0.4 1 0] vc +f +0.4 w +2 J +2 M +S +n +1982.7 2080.6 m +1981.5 2079.5 1980.5 2078.4 1979.3 2077.2 C +1980.5 2078.4 1981.5 2079.5 1982.7 2080.6 C +f +S +n +1988 2078.2 m +1989.4 2077.5 1990.8 2076.8 1992.3 2076 C +1990.8 2076.8 1989.4 2077.5 1988 2078.2 C +[0 0.2 1 0] vc +f +S +n +1988.7 2078.2 m +1991.1 2076.4 1991.8 2076.4 1988.7 2078.2 C +[0 0.4 1 0] vc +f +S +n +1976.2 2063.8 m +1978.6 2062.2 1976 2063.3 1976 2064.5 C +1976.1 2067.8 1975.5 2071.4 1976.4 2074.4 C +1975.7 2071.1 1975.9 2067.2 1976.2 2063.8 C +f +S +n +1996.8 2074.1 m +1998.3 2073.4 1999.7 2072.7 2001.2 2072 C +1999.7 2072.7 1998.3 2073.4 1996.8 2074.1 C +f +S +n +2001.6 2071.7 m +2002.9 2071.2 2004.2 2070.6 2005.5 2070 C +2004.2 2070.6 2002.9 2071.2 2001.6 2071.7 C +f +S +n +1981.5 2060.7 m +1980.2 2061.2 1978.9 2061.5 1977.9 2062.6 C +1978.9 2061.5 1980.2 2061.2 1981.5 2060.7 C +f +S +n +1982 2060.4 m +1982.7 2060.1 1983.6 2059.8 1984.4 2059.5 C +1983.6 2059.8 1982.7 2060.1 1982 2060.4 C +f +S +n +1952 2051.3 m +1950.8 2050.2 1949.7 2049.1 1948.6 2048 C +1949.7 2049.1 1950.8 2050.2 1952 2051.3 C +f +S +n +vmrs +1977.4 2047.7 m +1975.8 2047.8 1974.8 2046.1 1974.5 2045.3 C +1974.9 2044.4 1976 2044.5 1976.7 2044.8 C +1977.9 2045 1977 2048.4 1979.3 2047.5 C +1979.9 2047.5 1980.8 2048.6 1979.8 2049.2 C +1978.2 2050.4 1980.8 2049.5 1980.3 2049.4 C +1981.4 2049.8 1980.3 2048.4 1980.3 2048 C +1979.8 2047.5 1979 2046.6 1978.4 2046.5 C +1977.3 2045.9 1977.2 2043.3 1975.2 2044.6 C +1974.7 2045.3 1973.6 2045 1973.3 2045.8 C +1975 2046.3 1975.8 2049.8 1978.1 2049.4 C +1978.4 2050.9 1978.7 2048.5 1977.9 2049.2 C +1977.7 2048.7 1977.2 2047.8 1977.4 2047.7 C +[0 0.5 0.5 0.2] vc +f +0.4 w +2 J +2 M +S +n +1957.2 2048.9 m +1958.7 2048.2 1960.1 2047.5 1961.6 2046.8 C +1960.1 2047.5 1958.7 2048.2 1957.2 2048.9 C +[0 0.2 1 0] vc +f +S +n +1958 2048.9 m +1960.4 2047.1 1961.1 2047.1 1958 2048.9 C +[0 0.4 1 0] vc +f +S +n +1966.1 2044.8 m +1967.6 2044.1 1969 2043.4 1970.4 2042.7 C +1969 2043.4 1967.6 2044.1 1966.1 2044.8 C +f +S +n +1970.9 2042.4 m +1972.2 2041.9 1973.5 2041.3 1974.8 2040.8 C +1973.5 2041.3 1972.2 2041.9 1970.9 2042.4 C +f +S +n +2012 2034.5 m +2010.4 2034.6 2009.3 2032.9 2009.1 2032.1 C +2009.4 2031 2010.3 2031.3 2011.2 2031.6 C +2012.5 2031.8 2011.6 2035.2 2013.9 2034.3 C +2014.4 2034.3 2015.4 2035.4 2014.4 2036 C +2012.7 2037.2 2015.3 2036.3 2014.8 2036.2 C +2015.9 2036.6 2014.8 2035.2 2014.8 2034.8 C +2014.4 2034.3 2013.6 2033.4 2012.9 2033.3 C +2011.5 2031 2009.3 2029.4 2007.4 2028 C +2007.5 2026.5 2007.3 2027.9 2007.2 2028.3 C +2007.9 2028.8 2008.7 2029.1 2009.3 2030 C +2009.6 2030.7 2009 2031.9 2008.4 2031.6 C +2006.7 2031 2007.7 2028 2005 2028.8 C +2004.8 2028.6 2004.3 2028.2 2003.8 2028.3 C +2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C +2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C +2012.7 2036.1 2011.8 2035 2012 2034.5 C +[0 0.5 0.5 0.2] vc +f +S +n +1981.2 2005.2 m +1979.7 2005.3 1978.6 2003.6 1978.4 2002.8 C +1978.7 2001.8 1979.6 2002.1 1980.5 2002.4 C +1981.8 2002.5 1980.9 2005.9 1983.2 2005 C +1983.7 2005.1 1984.7 2006.1 1983.6 2006.7 C +1982 2007.9 1984.6 2007 1984.1 2006.9 C +1985.2 2007.3 1984.1 2006 1984.1 2005.5 C +1983.6 2005 1982.9 2004.1 1982.2 2004 C +1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C +1976.7 1997.2 1976.6 1998.6 1976.4 1999 C +1977.2 1999.5 1978 1999.8 1978.6 2000.7 C +1978.8 2001.5 1978.3 2002.7 1977.6 2002.4 C +1976 2001.8 1977 1998.7 1974.3 1999.5 C +1974.1 1999.3 1973.6 1998.9 1973.1 1999 C +1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C +1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C +1982 2006.8 1981.1 2005.7 1981.2 2005.2 C +f +S +n +1966.8 1976.4 m +1969.4 1973 1974.4 1974.6 1976.2 1970.4 C +1972.7 1974 1968 1975.1 1964 1977.4 C +1960.9 1979.9 1957.1 1981.8 1953.9 1982.7 C +1958.4 1981.1 1962.6 1978.8 1966.8 1976.4 C +[0.18 0.18 0 0.78] vc +f +S +n +1948.4 2093.8 m +1949.8 2093.1 1951.2 2092.5 1952.7 2091.9 C +1951.2 2092.5 1949.8 2093.1 1948.4 2093.8 C +[0 0.2 1 0] vc +f +S +n +1948.1 2093.6 m +1947.3 2092.8 1946.5 2091.9 1945.7 2091.2 C +1946.5 2091.9 1947.3 2092.8 1948.1 2093.6 C +f +S +n +vmrs +1942.1 2087.8 m +1943.5 2088.4 1944.3 2089.5 1945.2 2090.7 C +1944.8 2089.3 1943.3 2088.3 1942.1 2087.8 C +[0 0.2 1 0] vc +f +0.4 w +2 J +2 M +S +n +1933.5 2078.4 m +1933.5 2078 1933.2 2079 1933.7 2079.4 C +1935 2080.4 1936.2 2081.3 1937.1 2082.8 C +1936.7 2080.7 1933.7 2080.7 1933.5 2078.4 C +f +S +n +1982.9 2080.6 m +1984.4 2079.9 1985.8 2079.3 1987.2 2078.7 C +1985.8 2079.3 1984.4 2079.9 1982.9 2080.6 C +f +S +n +1982.7 2080.4 m +1981.9 2079.6 1981.1 2078.7 1980.3 2078 C +1981.1 2078.7 1981.9 2079.6 1982.7 2080.4 C +f +S +n +1977.4 2075.1 m +1977.9 2075.3 1979.1 2076.4 1979.8 2077.5 C +1979 2076.8 1978.7 2075.1 1977.4 2075.1 C +f +S +n +1952.2 2051.3 m +1953.6 2050.7 1955.1 2050.1 1956.5 2049.4 C +1955.1 2050.1 1953.6 2050.7 1952.2 2051.3 C +f +S +n +1952 2051.1 m +1951.2 2050.3 1950.3 2049.5 1949.6 2048.7 C +1950.3 2049.5 1951.2 2050.3 1952 2051.1 C +f +S +n +1946 2045.3 m +1947.3 2045.9 1948.1 2047 1949.1 2048.2 C +1948.6 2046.8 1947.1 2045.8 1946 2045.3 C +f +S +n +1937.3 2036 m +1937.4 2035.5 1937 2036.5 1937.6 2036.9 C +1938.8 2037.9 1940.1 2038.8 1940.9 2040.3 C +1940.6 2038.2 1937.6 2038.2 1937.3 2036 C +f +S +n +1935.2 2073.2 m +1936.4 2069.9 1935.8 2061.8 1935.6 2056.4 C +1935.8 2055.9 1936.3 2055.7 1936.1 2055.2 C +1935.7 2054.7 1935 2055 1934.4 2054.9 C +1934.4 2061.5 1934.4 2068.7 1934.4 2074.6 C +1935.7 2075.1 1936 2073.7 1935.2 2073.2 C +[0 0.01 1 0] vc +f +S +n +vmrs +1939 2030.7 m +1940.3 2027.4 1939.7 2019.3 1939.5 2013.9 C +1939.7 2013.5 1940.1 2013.2 1940 2012.7 C +1939.5 2012.3 1938.8 2012.5 1938.3 2012.4 C +1938.3 2019 1938.3 2026.2 1938.3 2032.1 C +1939.5 2032.7 1939.8 2031.2 1939 2030.7 C +[0 0.01 1 0] vc +f +0.4 w +2 J +2 M +S +n +1975.2 2077.2 m +1975.3 2077.3 1975.4 2077.4 1975.5 2077.5 C +1974.7 2073.2 1974.9 2067.5 1975.2 2063.6 C +1975.4 2064 1974.6 2063.9 1974.8 2064.3 C +1974.9 2069.9 1974.3 2076.5 1975.2 2081.1 C +1974.9 2079.9 1974.9 2078.4 1975.2 2077.2 C +[0.92 0.92 0 0.67] vc +f +S +n +1930.8 2067.4 m +1931.5 2070.1 1929.6 2072.1 1930.6 2074.6 C +1931 2072.6 1930.8 2069.8 1930.8 2067.4 C +f +S +n +2010 2050.1 m +2009.8 2050.5 2009.5 2050.9 2009.3 2051.1 C +2009.5 2056.7 2008.9 2063.3 2009.8 2067.9 C +2009.5 2062.1 2009.3 2054.7 2010 2050.1 C +f +S +n +1930.1 2060.9 m +1929.3 2057.1 1930.7 2054.8 1929.9 2051.3 C +1930.2 2050.2 1931.1 2049.6 1931.8 2049.2 C +1931.4 2049.6 1930.4 2049.5 1930.1 2050.1 C +1928.4 2054.8 1933.4 2063.5 1925.3 2064.3 C +1927.2 2063.9 1928.5 2062.1 1930.1 2060.9 C +[0.07 0.06 0 0.58] vc +f +S +n +1929.6 2061.2 m +1929.6 2057.6 1929.6 2054.1 1929.6 2050.6 C +1930 2049.9 1930.5 2049.4 1931.1 2049.2 C +1930 2048.6 1930.5 2050.2 1929.4 2049.6 C +1928 2054.4 1932.8 2063 1925.3 2064 C +1926.9 2063.3 1928.3 2062.4 1929.6 2061.2 C +[0.4 0.4 0 0] vc +f +S +n +1930.8 2061.6 m +1930.5 2058 1931.6 2054 1930.8 2051.3 C +1930.3 2054.5 1930.9 2058.5 1930.4 2061.9 C +1930.5 2061.2 1931 2062.2 1930.8 2061.6 C +[0.92 0.92 0 0.67] vc +f +S +n +1941.2 2045.1 m +1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C +1934.2 2040 1933.7 2036.4 1934 2039.3 C +1934.9 2040.1 1936.1 2039.9 1936.8 2040.8 C +1935.3 2044.2 1942.3 2041.7 1939.5 2046 C +1937.1 2048.5 1940.5 2045.6 1941.2 2045.1 C +f +S +n +1910 2045.8 m +1910 2039.4 1910 2033 1910 2026.6 C +1910 2033 1910 2039.4 1910 2045.8 C +f +S +n +1978.8 2022.3 m +1979.1 2021.7 1979.4 2020.4 1978.6 2021.6 C +1978.6 2026.9 1978.6 2033 1978.6 2037.6 C +1979.2 2037 1979.1 2038.2 1979.1 2038.6 C +1978.7 2033.6 1978.9 2026.8 1978.8 2022.3 C +f +S +n +vmrs +2026.1 2041.2 m +2026.1 2034.8 2026.1 2028.3 2026.1 2021.8 C +2026.1 2028.5 2026.3 2035.4 2025.9 2042 C +2024.4 2042.9 2022.9 2044.1 2021.3 2044.8 C +2023.1 2044 2025.1 2042.8 2026.1 2041.2 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +2026.4 2021.8 m +2026.3 2028.5 2026.5 2035.4 2026.1 2042 C +2025.6 2042.8 2024.7 2042.7 2024.2 2043.4 C +2024.7 2042.7 2025.5 2042.7 2026.1 2042.2 C +2026.5 2035.5 2026.3 2027.9 2026.4 2021.8 C +[0.4 0.4 0 0] vc +f +S +n +2025.6 2038.4 m +2025.6 2033 2025.6 2027.6 2025.6 2022.3 C +2025.6 2027.6 2025.6 2033 2025.6 2038.4 C +[0.92 0.92 0 0.67] vc +f +S +n +1934 2023.5 m +1934 2024.7 1933.8 2026 1934.2 2027.1 C +1934 2025.5 1934.7 2024.6 1934 2023.5 C +f +S +n +1928.2 2023.5 m +1928 2024.6 1927.4 2023.1 1926.8 2023.2 C +1926.2 2021 1921.4 2019.3 1923.2 2018 C +1922.7 2016.5 1923.2 2019.3 1922.2 2018.2 C +1924.4 2020.4 1926.2 2023.3 1928.9 2024.9 C +1927.9 2024.2 1929.8 2023.5 1928.2 2023.5 C +[0.18 0.18 0 0.78] vc +f +S +n +1934 2019.2 m +1932 2019.6 1930.8 2022.6 1928.7 2021.8 C +1924.5 2016.5 1918.2 2011.8 1914 2006.7 C +1914 2005.7 1914 2004.6 1914 2003.6 C +1913.6 2004.3 1913.9 2005.8 1913.8 2006.9 C +1919 2012.4 1924.1 2016.5 1929.2 2022.3 C +1931 2021.7 1932.2 2019.8 1934 2019.2 C +f +S +n +1928.7 2024.9 m +1926.3 2022.7 1924.1 2020.4 1921.7 2018.2 C +1924.1 2020.4 1926.3 2022.7 1928.7 2024.9 C +[0.65 0.65 0 0.42] vc +f +S +n +1914.3 2006.7 m +1918.7 2011.8 1924.5 2016.4 1928.9 2021.6 C +1924.2 2016.1 1919 2012.1 1914.3 2006.7 C +[0.07 0.06 0 0.58] vc +f +S +n +1924.8 2020.8 m +1921.2 2016.9 1925.6 2022.5 1926 2021.1 C +1924.2 2021 1926.7 2019.6 1924.8 2020.8 C +[0.92 0.92 0 0.67] vc +f +S +n +1934 2018.4 m +1933.2 2014.7 1934.5 2012.3 1933.7 2008.8 C +1934 2007.8 1935 2007.2 1935.6 2006.7 C +1935.3 2007.1 1934.3 2007 1934 2007.6 C +1932.2 2012.3 1937.2 2021 1929.2 2021.8 C +1931.1 2021.4 1932.3 2019.6 1934 2018.4 C +[0.07 0.06 0 0.58] vc +f +S +n +vmrs +1933.5 2018.7 m +1933.5 2015.1 1933.5 2011.7 1933.5 2008.1 C +1933.8 2007.4 1934.3 2006.9 1934.9 2006.7 C +1933.8 2006.1 1934.3 2007.7 1933.2 2007.2 C +1931.9 2012 1936.7 2020.5 1929.2 2021.6 C +1930.7 2020.8 1932.2 2019.9 1933.5 2018.7 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1934.7 2019.2 m +1934.3 2015.6 1935.4 2011.5 1934.7 2008.8 C +1934.1 2012 1934.7 2016 1934.2 2019.4 C +1934.4 2018.7 1934.8 2019.8 1934.7 2019.2 C +[0.92 0.92 0 0.67] vc +f +S +n +1917.6 2013.6 m +1917.8 2011.1 1916.8 2014.2 1917.2 2012.2 C +1916.3 2012.9 1914.8 2011.8 1914.3 2010.8 C +1914.2 2010.5 1914.4 2010.4 1914.5 2010.3 C +1913.9 2008.8 1913.9 2011.9 1914.3 2012 C +1916.3 2012 1917.6 2013.6 1916.7 2015.6 C +1913.7 2017.4 1919.6 2014.8 1917.6 2013.6 C +f +S +n +1887.2 2015.3 m +1887.2 2008.9 1887.2 2002.5 1887.2 1996.1 C +1887.2 2002.5 1887.2 2008.9 1887.2 2015.3 C +f +S +n +1916.7 2014.4 m +1917 2012.1 1913 2013 1913.8 2010.8 C +1912.1 2009.8 1910.9 2009.4 1910.7 2007.9 C +1910.4 2010.6 1913.4 2010.4 1914 2012.4 C +1914.9 2012.8 1916.6 2012.9 1916.4 2014.4 C +1916.9 2015.1 1914.5 2016.6 1916.2 2015.8 C +1916.4 2015.3 1916.7 2015 1916.7 2014.4 C +[0.65 0.65 0 0.42] vc +f +S +n +1914 2009.3 m +1912.8 2010.9 1909.6 2005.3 1911.9 2009.8 C +1912.3 2009.6 1913.6 2010.2 1914 2009.3 C +[0.92 0.92 0 0.67] vc +f +S +n +1951.2 1998.8 m +1949 1996.4 1951.5 1994 1950.3 1991.8 C +1949.1 1989.1 1954 1982.7 1948.8 1981.2 C +1949.2 1981.5 1951 1982.4 1950.8 1983.6 C +1951.9 1988.6 1947.1 1986.5 1948.1 1990.4 C +1948.5 1990.3 1948.7 1990.7 1948.6 1991.1 C +1949 1992.5 1947.3 1991.9 1948.1 1992.5 C +1947.1 1992.7 1945.7 1993.5 1945.2 1994.7 C +1944.5 1996.8 1947.7 2000.5 1943.8 2001.4 C +1943.4 2002 1943.7 2004 1942.4 2004.5 C +1945.2 2002.2 1948.9 2000.9 1951.2 1998.8 C +f +S +n +1994.9 1993 m +1995.1 1996.5 1994.5 2000.3 1995.4 2003.6 C +1994.5 2000.3 1995.1 1996.5 1994.9 1993 C +f +S +n +1913.8 2003.3 m +1913.8 1996.9 1913.8 1990.5 1913.8 1984.1 C +1913.8 1990.5 1913.8 1996.9 1913.8 2003.3 C +f +S +n +1941.9 1998 m +1940.5 1997.3 1940.7 1999.4 1940.7 2000 C +1942.8 2001.3 1942.6 1998.8 1941.9 1998 C +[0 0 0 0] vc +f +S +n +vmrs +1942.1 1999.2 m +1942.2 1998.9 1941.8 1998.8 1941.6 1998.5 C +1940.4 1998 1940.7 1999.7 1940.7 2000 C +1941.6 2000.3 1942.6 2000.4 1942.1 1999.2 C +[0.92 0.92 0 0.67] vc +f +0.4 w +2 J +2 M +S +n +1940 1997.1 m +1939.8 1996 1939.7 1995.9 1939.2 1995.2 C +1939.1 1995.3 1938.5 1997.9 1937.8 1996.4 C +1938 1997.3 1939.4 1998.6 1940 1997.1 C +f +S +n +1911.2 1995.9 m +1911.2 1991.6 1911.3 1987.2 1911.4 1982.9 C +1911.3 1987.2 1911.2 1991.6 1911.2 1995.9 C +f +S +n +1947.2 1979.1 m +1945.1 1978.8 1944.6 1975.7 1942.4 1975 C +1940.5 1972.6 1942.2 1973.7 1942.4 1975.7 C +1945.8 1975.5 1944.2 1979.8 1947.6 1979.6 C +1948.3 1982.3 1948.5 1980 1947.2 1979.1 C +f +S +n +1939.5 1973.3 m +1940.1 1972.6 1939.8 1974.2 1940.2 1973.1 C +1939.1 1972.8 1938.8 1968.5 1935.9 1969.7 C +1937.4 1969.2 1938.5 1970.6 1939 1971.4 C +1939.2 1972.7 1938.6 1973.9 1939.5 1973.3 C +f +S +n +1975.2 2073.2 m +1975.2 2070.2 1975.2 2067.2 1975.2 2064.3 C +1975.2 2067.2 1975.2 2070.2 1975.2 2073.2 C +[0.18 0.18 0 0.78] vc +f +S +n +1929.9 2065.7 m +1928.1 2065.6 1926 2068.8 1924.1 2066.9 C +1918.1 2060.9 1912.9 2055.7 1907.1 2049.9 C +1906.7 2047.1 1906.9 2043.9 1906.8 2041 C +1906.8 2043.9 1906.8 2046.8 1906.8 2049.6 C +1913.2 2055.5 1918.7 2061.9 1925.1 2067.6 C +1927.1 2067.9 1928.6 2064.4 1930.1 2066.2 C +1929.7 2070.3 1929.9 2074.7 1929.9 2078.9 C +1929.6 2074.4 1930.5 2070.1 1929.9 2065.7 C +[0.07 0.06 0 0.58] vc +f +S +n +1930.1 2061.6 m +1928.1 2062.1 1927 2065.1 1924.8 2064.3 C +1920.7 2058.9 1914.4 2054.3 1910.2 2049.2 C +1910.2 2048.1 1910.2 2047.1 1910.2 2046 C +1909.8 2046.8 1910 2048.3 1910 2049.4 C +1915.1 2054.9 1920.3 2059 1925.3 2064.8 C +1927.1 2064.2 1928.4 2062.3 1930.1 2061.6 C +[0.18 0.18 0 0.78] vc +f +S +n +1932 2049.9 m +1932.3 2050.3 1932 2050.4 1932.8 2050.4 C +1932 2050.4 1932.2 2049.2 1931.3 2049.6 C +1931.4 2050.5 1930.3 2050.4 1930.4 2051.3 C +1931.1 2051.1 1930.7 2049.4 1932 2049.9 C +f +S +n +1938.3 2046 m +1936.3 2046.8 1935.2 2047.2 1934.2 2048.9 C +1935.3 2047.7 1936.8 2046.2 1938.3 2046 C +[0.4 0.4 0 0] vc +f +S +n +vmrs +1938.3 2047 m +1937.9 2046.9 1936.6 2047.1 1936.1 2048 C +1936.5 2047.5 1937.3 2046.7 1938.3 2047 C +[0.18 0.18 0 0.78] vc +f +0.4 w +2 J +2 M +S +n +1910.2 2043.2 m +1910.1 2037.5 1910 2031.8 1910 2026.1 C +1910 2031.8 1910.1 2037.5 1910.2 2043.2 C +f +S +n +1933.5 2032.1 m +1933.7 2035.2 1932.8 2035.8 1933.7 2038.6 C +1933.3 2036.6 1934.6 2018 1933.5 2032.1 C +f +S +n +1907.3 2021.8 m +1906.6 2025.9 1909.4 2032.6 1903.2 2034 C +1902.8 2034.1 1902.4 2033.9 1902 2033.8 C +1897.9 2028.5 1891.6 2023.8 1887.4 2018.7 C +1887.4 2017.7 1887.4 2016.6 1887.4 2015.6 C +1887 2016.3 1887.2 2017.8 1887.2 2018.9 C +1892.3 2024.4 1897.5 2028.5 1902.5 2034.3 C +1904.3 2033.6 1905.7 2032 1907.3 2030.9 C +1907.3 2027.9 1907.3 2024.9 1907.3 2021.8 C +f +S +n +1933.7 2023.2 m +1932 2021.7 1931.1 2024.9 1929.4 2024.9 C +1931.2 2024.7 1932.4 2021.5 1933.7 2023.2 C +f +S +n +1989.2 2024.4 m +1987.4 2023.7 1985.8 2022.2 1985.1 2020.4 C +1984.6 2020.1 1986 2018.9 1985.1 2019.2 C +1985.6 2020.8 1984.1 2019.4 1984.6 2021.1 C +1986.3 2022.3 1988.1 2025.3 1989.2 2024.4 C +f +S +n +1904.4 2031.9 m +1903 2029.7 1905.3 2027.7 1904.2 2025.9 C +1904.5 2025 1903.7 2023 1904 2021.3 C +1904 2022.3 1903.2 2022 1902.5 2022 C +1901.3 2022.3 1902.2 2020.1 1901.6 2019.6 C +1902.5 2019.8 1902.6 2018.3 1903.5 2018.9 C +1903.7 2021.8 1905.6 2016.8 1905.6 2020.6 C +1905.9 2020 1906.3 2020.8 1906.1 2021.1 C +1905.8 2022.7 1906.7 2020.4 1906.4 2019.9 C +1906.4 2018.5 1908.2 2017.8 1906.8 2016.5 C +1906.9 2015.7 1907.7 2017.1 1907.1 2016.3 C +1908.5 2015.8 1910.3 2015.1 1911.6 2016 C +1912.2 2016.2 1911.9 2018 1911.6 2018 C +1914.5 2017.1 1910.4 2013.6 1913.3 2013.4 C +1912.4 2011.3 1910.5 2011.8 1909.5 2010 C +1910 2010.5 1909 2010.8 1908.8 2011.2 C +1907.5 2009.9 1906.1 2011.7 1904.9 2011.5 C +1904.7 2010.9 1904.3 2010.5 1904.4 2009.8 C +1905 2010.2 1904.6 2008.6 1905.4 2008.1 C +1906.6 2007.5 1907.7 2008.4 1908.5 2007.4 C +1908.9 2008.5 1909.7 2008.1 1909 2007.2 C +1908.1 2006.5 1905.9 2007.3 1905.4 2007.4 C +1903.9 2007.3 1905.2 2008.5 1904.2 2008.4 C +1904.6 2009.9 1902.8 2010.3 1902.3 2010.5 C +1901.5 2009.9 1900.4 2010 1899.4 2010 C +1898.6 2011.2 1898.2 2013.4 1896.5 2013.4 C +1896 2012.9 1894.4 2012.9 1893.6 2012.9 C +1893.1 2013.9 1892.9 2015.5 1891.5 2016 C +1890.3 2016.1 1889.2 2014 1888.6 2015.8 C +1890 2016 1891 2016.9 1892.9 2016.5 C +1894.1 2017.2 1892.8 2018.3 1893.2 2018.9 C +1892.6 2018.9 1891.1 2019.8 1890.5 2020.6 C +1891.1 2023.6 1893.2 2019.8 1893.9 2022.5 C +1894.1 2023.3 1892.7 2023.6 1893.9 2024 C +1894.2 2024.3 1897.4 2023.8 1896.5 2026.1 C +1896 2025.6 1897.4 2028.1 1897.5 2027.1 C +1898.4 2027.4 1899.3 2027 1899.6 2028.5 C +1899.5 2028.6 1899.4 2028.8 1899.2 2028.8 C +1899.3 2029.2 1899.6 2029.8 1900.1 2030.2 C +1900.4 2029.6 1901 2030 1901.8 2030.2 C +1903.1 2032.1 1900.4 2031.5 1902.8 2033.1 C +1903.3 2032.7 1904.5 2032 1904.4 2031.9 C +[0.21 0.21 0 0] vc +f +S +n +1909.2 2019.4 m +1908.8 2020.3 1910.2 2019.8 1909.2 2019.2 C +1908.3 2019.3 1907.6 2020.2 1907.6 2021.3 C +1908.5 2021 1907.6 2019 1909.2 2019.4 C +[0.18 0.18 0 0.78] vc +f +S +n +1915.5 2015.6 m +1913.5 2016.3 1912.4 2016.8 1911.4 2018.4 C +1912.5 2017.2 1914 2015.7 1915.5 2015.6 C +[0.4 0.4 0 0] vc +f +S +n +1915.5 2016.5 m +1915.1 2016.4 1913.8 2016.6 1913.3 2017.5 C +1913.7 2017 1914.5 2016.2 1915.5 2016.5 C +[0.18 0.18 0 0.78] vc +f +S +n +vmrs +1887.4 2012.7 m +1887.3 2007 1887.2 2001.3 1887.2 1995.6 C +1887.2 2001.3 1887.3 2007 1887.4 2012.7 C +[0.18 0.18 0 0.78] vc +f +0.4 w +2 J +2 M +S +n +1935.9 2007.4 m +1936.2 2007.8 1935.8 2007.9 1936.6 2007.9 C +1935.9 2007.9 1936.1 2006.7 1935.2 2007.2 C +1935.2 2008.1 1934.1 2007.9 1934.2 2008.8 C +1935 2008.7 1934.6 2006.9 1935.9 2007.4 C +f +S +n +1942.1 2003.6 m +1940.1 2004.3 1939.1 2004.8 1938 2006.4 C +1939.1 2005.2 1940.6 2003.7 1942.1 2003.6 C +[0.4 0.4 0 0] vc +f +S +n +1942.1 2004.5 m +1941.8 2004.4 1940.4 2004.6 1940 2005.5 C +1940.4 2005 1941.2 2004.2 1942.1 2004.5 C +[0.18 0.18 0 0.78] vc +f +S +n +1914 2000.7 m +1914 1995 1913.9 1989.3 1913.8 1983.6 C +1913.9 1989.3 1914 1995 1914 2000.7 C +f +S +n +1941.6 1998.3 m +1943.4 2001.9 1942.4 1996 1940.9 1998.3 C +1941.2 1998.3 1941.4 1998.3 1941.6 1998.3 C +f +S +n +1954.8 1989.9 m +1953.9 1989.6 1954.7 1991.6 1953.9 1991.1 C +1954.5 1993.1 1953.6 1998 1954.6 1993.2 C +1954 1992.2 1954.7 1990.7 1954.8 1989.9 C +f +S +n +1947.6 1992.5 m +1946.2 1993.5 1944.9 1993 1944.8 1994.7 C +1945.5 1994 1947 1992.2 1947.6 1992.5 C +f +S +n +1910.7 1982.2 m +1910.3 1981.8 1909.7 1982 1909.2 1982 C +1909.7 1982 1910.3 1981.9 1910.7 1982.2 C +1911 1987.1 1910 1992.6 1910.7 1997.3 C +1910.7 1992.3 1910.7 1987.2 1910.7 1982.2 C +[0.65 0.65 0 0.42] vc +f +S +n +1910.9 1992.8 m +1910.9 1991.3 1910.9 1989.7 1910.9 1988.2 C +1910.9 1989.7 1910.9 1991.3 1910.9 1992.8 C +[0.18 0.18 0 0.78] vc +f +S +n +vmrs +1953.6 1983.6 m +1954.1 1985.3 1953.2 1988.6 1954.8 1989.4 C +1954.1 1987.9 1954.4 1985.4 1953.6 1983.6 C +[0.18 0.18 0 0.78] vc +f +0.4 w +2 J +2 M +S +n +1910.7 1982 m +1911.6 1982.9 1911 1984.4 1911.2 1985.6 C +1911 1984.4 1911.6 1982.9 1910.7 1982 C +f +S +n +1947.2 1979.6 m +1947.5 1980.6 1948.3 1980.6 1947.4 1979.6 C +1946.2 1979.4 1945.7 1978.8 1947.2 1979.6 C +f +S +n +1930.4 2061.4 m +1930.4 2058 1930.4 2053.5 1930.4 2051.1 C +1930.7 2054.6 1929.8 2057.4 1930.1 2061.2 C +1929.5 2061.9 1929.7 2061.2 1930.4 2061.4 C +[0.65 0.65 0 0.42] vc +f +S +n +1939.5 2044.8 m +1940 2041.5 1935.2 2044.3 1936.4 2040.8 C +1934.9 2040.9 1934.1 2039.7 1933.5 2038.6 C +1933.3 2035.4 1933.2 2040 1934 2040.3 C +1936.2 2040.6 1936.3 2043.6 1938.5 2043.4 C +1939.7 2044.2 1939.4 2045.6 1938.3 2046.5 C +1939.1 2046.6 1939.6 2045.6 1939.5 2044.8 C +f +S +n +1910.4 2045.3 m +1910.4 2039.5 1910.4 2033.6 1910.4 2027.8 C +1910.4 2033.6 1910.4 2039.5 1910.4 2045.3 C +f +S +n +1906.8 2030.9 m +1907.6 2026.8 1905 2020.8 1909 2018.7 C +1906.5 2018.9 1906.8 2022.4 1906.8 2024.7 C +1906.4 2028.2 1907.9 2032 1903 2033.8 C +1902.2 2034 1903.8 2033.4 1904.2 2033.1 C +1905.1 2032.4 1905.9 2031.5 1906.8 2030.9 C +[0.07 0.06 0 0.58] vc +f +S +n +1907.1 2030.7 m +1907.1 2028.8 1907.1 2027 1907.1 2025.2 C +1907.1 2027 1907.1 2028.8 1907.1 2030.7 C +[0.65 0.65 0 0.42] vc +f +S +n +1932 2023.2 m +1932.2 2023.6 1931.7 2023.7 1931.6 2024 C +1932 2023.7 1932.3 2022.8 1933 2023 C +1933.9 2024.3 1933.3 2026.2 1933.5 2027.8 C +1933.5 2026.4 1934.9 2022.2 1932 2023.2 C +f +S +n +2026.1 2021.6 m +2026.1 2020.8 2026.1 2019.9 2026.1 2019.2 C +2026.1 2019.9 2026.1 2020.8 2026.1 2021.6 C +f +S +n +vmrs +1934.2 2018.9 m +1934.2 2015.5 1934.2 2011 1934.2 2008.6 C +1934.5 2012.1 1933.7 2014.9 1934 2018.7 C +1933.4 2019.5 1933.5 2018.7 1934.2 2018.9 C +[0.65 0.65 0 0.42] vc +f +0.4 w +2 J +2 M +S +n +1887.6 2014.8 m +1887.6 2009 1887.6 2003.1 1887.6 1997.3 C +1887.6 2003.1 1887.6 2009 1887.6 2014.8 C +f +S +n +1914.3 2002.8 m +1914.3 1997 1914.3 1991.1 1914.3 1985.3 C +1914.3 1991.1 1914.3 1997 1914.3 2002.8 C +f +S +n +1995.4 1992.3 m +1995.4 1991.5 1995.4 1990.7 1995.4 1989.9 C +1995.4 1990.7 1995.4 1991.5 1995.4 1992.3 C +f +S +n +1896 1988.4 m +1896.9 1988 1897.8 1987.7 1898.7 1987.2 C +1897.8 1987.7 1896.9 1988 1896 1988.4 C +f +S +n +1899.4 1986.8 m +1900.4 1986.3 1901.3 1985.8 1902.3 1985.3 C +1901.3 1985.8 1900.4 1986.3 1899.4 1986.8 C +f +S +n +1902.8 1985.1 m +1905.2 1984 1905.2 1984 1902.8 1985.1 C +f +S +n +1949.1 1983.4 m +1950.2 1984.4 1947.8 1984.6 1949.3 1985.1 C +1949.5 1984.4 1949.6 1984.1 1949.1 1983.4 C +[0.07 0.06 0 0.58] vc +f +S +n +1906.1 1983.4 m +1908.6 1982 1908.6 1982 1906.1 1983.4 C +[0.65 0.65 0 0.42] vc +f +S +n +1922.7 1976.4 m +1923.6 1976 1924.4 1975.7 1925.3 1975.2 C +1924.4 1975.7 1923.6 1976 1922.7 1976.4 C +f +S +n +vmrs +1926 1974.8 m +1927 1974.3 1928 1973.8 1928.9 1973.3 C +1928 1973.8 1927 1974.3 1926 1974.8 C +[0.65 0.65 0 0.42] vc +f +0.4 w +2 J +2 M +S +n +1929.4 1973.1 m +1931.9 1972 1931.9 1972 1929.4 1973.1 C +f +S +n +1932.8 1971.4 m +1935.3 1970 1935.3 1970 1932.8 1971.4 C +f +S +n +1949.6 2097.2 m +1951.1 2096.4 1952.6 2095.5 1954.1 2094.8 C +1952.6 2095.5 1951.1 2096.4 1949.6 2097.2 C +[0.07 0.06 0 0.58] vc +f +S +n +1955.1 2094.3 m +1956.7 2093.5 1958.3 2092.7 1959.9 2091.9 C +1958.3 2092.7 1956.7 2093.5 1955.1 2094.3 C +f +S +n +1960.4 2091.6 m +1961.3 2091.2 1962.1 2090.9 1963 2090.4 C +1962.1 2090.9 1961.3 2091.2 1960.4 2091.6 C +f +S +n +1963.5 2090.2 m +1964.4 2089.7 1965.2 2089.2 1966.1 2088.8 C +1965.2 2089.2 1964.4 2089.7 1963.5 2090.2 C +f +S +n +1966.6 2088.5 m +1969.5 2087.1 1972.4 2085.8 1975.2 2084.4 C +1972.4 2085.8 1969.5 2087.1 1966.6 2088.5 C +f +S +n +1965.2 2086.1 m +1965.9 2085.7 1966.8 2085.3 1967.6 2084.9 C +1966.8 2085.3 1965.9 2085.7 1965.2 2086.1 C +f +S +n +1968.3 2084.7 m +1969.2 2084.3 1970 2083.9 1970.9 2083.5 C +1970 2083.9 1969.2 2084.3 1968.3 2084.7 C +f +S +n +vmrs +1984.1 2084 m +1985.6 2083.2 1987.2 2082.3 1988.7 2081.6 C +1987.2 2082.3 1985.6 2083.2 1984.1 2084 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1976 2078.7 m +1978.1 2080.1 1980 2082 1982 2083.7 C +1980 2081.9 1977.9 2080.3 1976 2078.2 C +1975.5 2079.9 1975.8 2081.9 1975.7 2083.7 C +1975.8 2082 1975.5 2080.2 1976 2078.7 C +f +S +n +1989.6 2081.1 m +1991.3 2080.3 1992.8 2079.5 1994.4 2078.7 C +1992.8 2079.5 1991.3 2080.3 1989.6 2081.1 C +f +S +n +1933.2 2074.6 m +1932.4 2076.2 1932.8 2077.5 1933 2078.7 C +1933 2077.6 1932.9 2074.8 1933.2 2074.6 C +f +S +n +1994.9 2078.4 m +1995.8 2078 1996.7 2077.7 1997.6 2077.2 C +1996.7 2077.7 1995.8 2078 1994.9 2078.4 C +f +S +n +1998 2077 m +1998.9 2076.5 1999.8 2076 2000.7 2075.6 C +1999.8 2076 1998.9 2076.5 1998 2077 C +f +S +n +2001.2 2075.3 m +2004 2073.9 2006.9 2072.6 2009.8 2071.2 C +2006.9 2072.6 2004 2073.9 2001.2 2075.3 C +f +S +n +1980.5 2060.7 m +1979.9 2060.7 1976.7 2062.8 1975.7 2064.5 C +1975.7 2067.5 1975.7 2070.5 1975.7 2073.4 C +1976.3 2068.7 1973.9 2061.6 1980.5 2060.7 C +f +S +n +1999.7 2072.9 m +2000.5 2072.5 2001.3 2072.1 2002.1 2071.7 C +2001.3 2072.1 2000.5 2072.5 1999.7 2072.9 C +f +S +n +2002.8 2071.5 m +2003.7 2071.1 2004.6 2070.7 2005.5 2070.3 C +2004.6 2070.7 2003.7 2071.1 2002.8 2071.5 C +f +S +n +vmrs +2015.1 2047.5 m +2014.4 2047.5 2011.2 2049.6 2010.3 2051.3 C +2010.3 2057.7 2010.3 2064.1 2010.3 2070.5 C +2010.3 2063.9 2010.1 2057.1 2010.5 2050.6 C +2012 2049.3 2013.5 2048.3 2015.1 2047.5 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1910.4 2049.2 m +1914.8 2054.3 1920.7 2058.9 1925.1 2064 C +1920.4 2058.6 1915.1 2054.6 1910.4 2049.2 C +f +S +n +1988.2 2057.3 m +1989.1 2056.8 1989.9 2056.2 1990.8 2055.6 C +1989.9 2056.2 1989.1 2056.8 1988.2 2057.3 C +f +S +n +1991.6 2051.3 m +1991.6 2046.3 1991.6 2041.2 1991.6 2036.2 C +1991.6 2041.2 1991.6 2046.3 1991.6 2051.3 C +f +S +n +1935.6 2047.5 m +1932.9 2051.7 1939.7 2043.8 1935.6 2047.5 C +f +S +n +1938.8 2043.9 m +1938.1 2043.3 1938.2 2043.7 1937.3 2043.4 C +1938.7 2043 1938.2 2044.9 1939 2045.3 C +1938.2 2045.3 1938.7 2046.6 1937.8 2046.5 C +1939.1 2046.2 1939.1 2044.5 1938.8 2043.9 C +f +S +n +1972.4 2045.6 m +1973.4 2045 1974.5 2044.4 1975.5 2043.9 C +1974.5 2044.4 1973.4 2045 1972.4 2045.6 C +f +S +n +1969 2043.6 m +1969.8 2043.2 1970.6 2042.9 1971.4 2042.4 C +1970.6 2042.9 1969.8 2043.2 1969 2043.6 C +f +S +n +1972.1 2042.2 m +1973 2041.8 1973.9 2041.4 1974.8 2041 C +1973.9 2041.4 1973 2041.8 1972.1 2042.2 C +f +S +n +1906.6 2035 m +1905 2034.7 1904.8 2036.6 1903.5 2036.9 C +1904.9 2037 1905.8 2033.4 1907.1 2035.7 C +1907.1 2037.2 1907.1 2038.6 1907.1 2040 C +1906.9 2038.4 1907.5 2036.4 1906.6 2035 C +f +S +n +vmrs +1937.1 2032.1 m +1936.2 2033.7 1936.6 2035 1936.8 2036.2 C +1936.8 2035.1 1936.8 2032.4 1937.1 2032.1 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1887.6 2018.7 m +1892 2023.8 1897.9 2028.4 1902.3 2033.6 C +1897.6 2028.1 1892.3 2024.1 1887.6 2018.7 C +f +S +n +1999.7 2031.4 m +1998.7 2030.3 1997.6 2029.2 1996.6 2028 C +1997.6 2029.2 1998.7 2030.3 1999.7 2031.4 C +f +S +n +1912.8 2017 m +1910.6 2021.1 1913.6 2015.3 1914.5 2016 C +1914 2016.3 1913.4 2016.7 1912.8 2017 C +f +S +n +1939.5 2005 m +1936.7 2009.2 1943.6 2001.3 1939.5 2005 C +f +S +n +1942.6 2001.4 m +1941.9 2000.8 1942 2001.2 1941.2 2000.9 C +1942.5 2000.6 1942.1 2002.4 1942.8 2002.8 C +1942 2002.8 1942.5 2004.1 1941.6 2004 C +1943 2003.7 1942.9 2002.1 1942.6 2001.4 C +f +S +n +2006.2 2000.7 m +2005.4 2001.5 2004 2002.8 2004 2002.8 C +2004.5 2002.4 2005.5 2001.4 2006.2 2000.7 C +f +S +n +1998.5 2001.6 m +1997.7 2002 1996.8 2002.4 1995.9 2002.6 C +1995.5 1999.3 1995.7 1995.7 1995.6 1992.3 C +1995.6 1995.7 1995.6 1999.2 1995.6 2002.6 C +1996.6 2002.4 1997.7 2002.2 1998.5 2001.6 C +[0.4 0.4 0 0] vc +f +S +n +1996.1 2002.8 m +1995.9 2002.8 1995.8 2002.8 1995.6 2002.8 C +1995.2 1999.5 1995.5 1995.9 1995.4 1992.5 C +1995.4 1995.9 1995.4 1999.4 1995.4 2002.8 C +1996.4 2003.1 1998.2 2001.6 1996.1 2002.8 C +[0.07 0.06 0 0.58] vc +f +S +n +1969 2002.1 m +1968 2001 1966.9 1999.9 1965.9 1998.8 C +1966.9 1999.9 1968 2001 1969 2002.1 C +f +S +n +vmrs +2000 2001.2 m +2002.1 2000 2004.1 1998.9 2006.2 1997.8 C +2004.1 1998.9 2002.1 2000 2000 2001.2 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1895.8 1984.8 m +1898.3 1983.6 1900.8 1982.3 1903.2 1981 C +1900.8 1982.3 1898.3 1983.6 1895.8 1984.8 C +f +S +n +1905.2 1980.3 m +1906.4 1979.9 1907.6 1979.5 1908.8 1979.1 C +1907.6 1979.5 1906.4 1979.9 1905.2 1980.3 C +f +S +n +1964.7 1977.4 m +1963.8 1977.5 1962.5 1980.2 1960.8 1980 C +1962.5 1980.2 1963.3 1978 1964.7 1977.4 C +f +S +n +1952 1979.6 m +1955.2 1979.2 1955.2 1979.2 1952 1979.6 C +f +S +n +1937.8 1966.4 m +1941.2 1969.5 1946.1 1976.4 1951.5 1979.3 C +1946.1 1976.7 1942.8 1970.4 1937.8 1966.4 C +f +S +n +1911.9 1978.6 m +1914.3 1977.4 1916.7 1976.2 1919.1 1975 C +1916.7 1976.2 1914.3 1977.4 1911.9 1978.6 C +f +S +n +1975.5 1971.4 m +1974.6 1972.2 1973.3 1973.6 1973.3 1973.6 C +1973.7 1973.1 1974.8 1972.1 1975.5 1971.4 C +f +S +n +1922.4 1972.8 m +1924.9 1971.6 1927.4 1970.3 1929.9 1969 C +1927.4 1970.3 1924.9 1971.6 1922.4 1972.8 C +f +S +n +1969.2 1971.9 m +1971.1 1970.9 1972.9 1969.8 1974.8 1968.8 C +1972.9 1969.8 1971.1 1970.9 1969.2 1971.9 C +f +S +n +vmrs +1931.8 1968.3 m +1933 1967.9 1934.2 1967.5 1935.4 1967.1 C +1934.2 1967.5 1933 1967.9 1931.8 1968.3 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1940.7 2072.4 m +1941.5 2072.4 1942.3 2072.3 1943.1 2072.2 C +1942.3 2072.3 1941.5 2072.4 1940.7 2072.4 C +[0 0 0 0.18] vc +f +S +n +1948.6 2069.3 m +1947 2069.5 1945.7 2068.9 1944.8 2069.8 C +1945.9 2068.5 1948.4 2070.2 1948.6 2069.3 C +f +S +n +1954.6 2066.4 m +1954.7 2067.9 1955.6 2067.3 1955.6 2068.8 C +1955.4 2067.8 1956 2066.6 1954.6 2066.4 C +f +S +n +1929.2 2061.2 m +1927.8 2062.1 1926.3 2064.1 1924.8 2063.3 C +1926.3 2064.6 1928 2062 1929.2 2061.2 C +f +S +n +1924.4 2067.4 m +1918.5 2061.6 1912.7 2055.9 1906.8 2050.1 C +1912.7 2055.9 1918.5 2061.6 1924.4 2067.4 C +[0.4 0.4 0 0] vc +f +S +n +1924.6 2062.8 m +1923.9 2062.1 1923.2 2061.2 1922.4 2060.4 C +1923.2 2061.2 1923.9 2062.1 1924.6 2062.8 C +[0 0 0 0.18] vc +f +S +n +1919.3 2057.3 m +1917.5 2055.6 1915.7 2053.8 1913.8 2052 C +1915.7 2053.8 1917.5 2055.6 1919.3 2057.3 C +f +S +n +1929.2 2055.2 m +1929.2 2054.2 1929.2 2053.2 1929.2 2052.3 C +1929.2 2053.2 1929.2 2054.2 1929.2 2055.2 C +f +S +n +1926.3 2049.6 m +1925.4 2049 1925.4 2050.5 1924.4 2050.4 C +1925.3 2051.3 1924.5 2051.9 1925.6 2052.5 C +1926.9 2052.6 1926 2050.6 1926.3 2049.6 C +f +S +n +vmrs +1911.2 2046.8 m +1910.1 2048.9 1911.9 2050.1 1913.1 2051.3 C +1912.1 2049.9 1910.6 2048.8 1911.2 2046.8 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1934 2048.7 m +1932.6 2048.7 1930.1 2047.7 1929.6 2049.4 C +1930.9 2048.6 1933.3 2049 1934 2048.7 C +f +S +n +1980 2048.4 m +1979.5 2046.8 1976.3 2047.9 1977.2 2045.6 C +1976.8 2045.1 1976.1 2044.7 1975.2 2044.8 C +1973.7 2046 1976.3 2046.4 1976.7 2047.5 C +1977.8 2047.2 1978.2 2050 1979.6 2049.2 C +1980 2049 1979.6 2048.6 1980 2048.4 C +f +S +n +1938.3 2045.6 m +1938.2 2044.4 1936.8 2043.8 1935.9 2043.4 C +1936.4 2044.4 1939.1 2044.3 1937.6 2045.8 C +1937 2046.1 1935.9 2046.1 1935.9 2046.8 C +1936.7 2046.3 1937.8 2046.2 1938.3 2045.6 C +f +S +n +1932.5 2040 m +1932.8 2038.1 1932 2038.9 1932.3 2040.3 C +1933.1 2040.3 1932.7 2041.7 1933.7 2041.5 C +1933.1 2041 1932.9 2040.5 1932.5 2040 C +f +S +n +2014.6 2035.2 m +2014.1 2033.6 2010.9 2034.7 2011.7 2032.4 C +2011.3 2031.9 2009.4 2030.7 2009.3 2032.1 C +2009.5 2033.7 2012.9 2033.8 2012.4 2035.7 C +2013 2036.4 2014.2 2036.5 2014.6 2035.2 C +f +S +n +1906.4 2030.7 m +1905 2031.6 1903.5 2033.6 1902 2032.8 C +1903.4 2034 1905.6 2031.4 1906.4 2030.7 C +f +S +n +1901.8 2037.2 m +1899.5 2034.8 1897.2 2032.5 1894.8 2030.2 C +1897.2 2032.5 1899.5 2034.8 1901.8 2037.2 C +[0.4 0.4 0 0] vc +f +S +n +1901.8 2032.4 m +1901.1 2031.6 1900.4 2030.7 1899.6 2030 C +1900.4 2030.7 1901.1 2031.6 1901.8 2032.4 C +[0 0 0 0.18] vc +f +S +n +1944.5 2030 m +1945.3 2029.9 1946.1 2029.8 1946.9 2029.7 C +1946.1 2029.8 1945.3 2029.9 1944.5 2030 C +f +S +n +vmrs +1997.8 2027.8 m +1997.7 2027.9 1997.6 2028.1 1997.3 2028 C +1997.4 2029.1 1998.5 2029.5 1999.2 2030 C +2000.1 2029.5 1998.9 2028 1997.8 2027.8 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1906.4 2029.2 m +1906.4 2026.6 1906.4 2024 1906.4 2021.3 C +1906.4 2024 1906.4 2026.6 1906.4 2029.2 C +f +S +n +2006.2 2025.9 m +2006 2025.9 2005.8 2025.8 2005.7 2025.6 C +2005.7 2025.5 2005.7 2025.3 2005.7 2025.2 C +2004.6 2025.8 2002.7 2024.7 2001.9 2026.1 C +2001.9 2027.9 2007.8 2029.2 2006.2 2025.9 C +[0 0 0 0] vc +f +S +n +1952.4 2026.8 m +1950.9 2027 1949.6 2026.4 1948.6 2027.3 C +1949.7 2026.1 1952.2 2027.7 1952.4 2026.8 C +[0 0 0 0.18] vc +f +S +n +1896.5 2026.8 m +1894.7 2025.1 1892.9 2023.3 1891 2021.6 C +1892.9 2023.3 1894.7 2025.1 1896.5 2026.8 C +f +S +n +1958.4 2024 m +1958.5 2025.5 1959.4 2024.8 1959.4 2026.4 C +1959.3 2025.3 1959.8 2024.1 1958.4 2024 C +f +S +n +1903.5 2019.2 m +1902.6 2018.6 1902.6 2020 1901.6 2019.9 C +1902.5 2020.8 1901.7 2021.4 1902.8 2022 C +1904.1 2022.2 1903.2 2020.1 1903.5 2019.2 C +f +S +n +1933 2018.7 m +1931.7 2019.6 1930.1 2021.6 1928.7 2020.8 C +1930.1 2022.1 1931.8 2019.5 1933 2018.7 C +f +S +n +1888.4 2016.3 m +1887.3 2018.4 1889.1 2019.6 1890.3 2020.8 C +1889.3 2019.5 1887.8 2018.3 1888.4 2016.3 C +f +S +n +1928.4 2020.4 m +1927.7 2019.6 1927 2018.7 1926.3 2018 C +1927 2018.7 1927.7 2019.6 1928.4 2020.4 C +f +S +n +vmrs +1911.2 2018.2 m +1909.8 2018.3 1907.3 2017.2 1906.8 2018.9 C +1908.1 2018.1 1910.5 2018.6 1911.2 2018.2 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1915.5 2015.1 m +1915.4 2013.9 1914 2013.3 1913.1 2012.9 C +1913.6 2013.9 1916.3 2013.8 1914.8 2015.3 C +1914.2 2015.6 1913.1 2015.6 1913.1 2016.3 C +1913.9 2015.9 1915 2015.7 1915.5 2015.1 C +f +S +n +1923.2 2014.8 m +1921.3 2013.1 1919.5 2011.3 1917.6 2009.6 C +1919.5 2011.3 1921.3 2013.1 1923.2 2014.8 C +f +S +n +1933 2012.7 m +1933 2011.7 1933 2010.8 1933 2009.8 C +1933 2010.8 1933 2011.7 1933 2012.7 C +f +S +n +1909.7 2008.1 m +1908.9 2009.2 1910.1 2009.9 1910.4 2011 C +1911.1 2010.7 1908.9 2009.7 1909.7 2008.1 C +f +S +n +1930.1 2007.2 m +1929.2 2006.6 1929.2 2008 1928.2 2007.9 C +1929.1 2008.8 1928.4 2009.4 1929.4 2010 C +1930.7 2010.2 1929.9 2008.1 1930.1 2007.2 C +f +S +n +1915 2004.3 m +1914 2006.4 1915.7 2007.6 1916.9 2008.8 C +1915.9 2007.5 1914.4 2006.3 1915 2004.3 C +f +S +n +1937.8 2006.2 m +1936.4 2006.3 1934 2005.2 1933.5 2006.9 C +1934.7 2006.1 1937.1 2006.6 1937.8 2006.2 C +f +S +n +1983.9 2006 m +1983.3 2004.3 1980.2 2005.4 1981 2003.1 C +1980.6 2002.7 1978.7 2001.5 1978.6 2002.8 C +1978.8 2004.4 1982.1 2004.5 1981.7 2006.4 C +1982.3 2007.2 1983.5 2007.2 1983.9 2006 C +f +S +n +1942.1 2003.1 m +1942 2001.9 1940.6 2001.3 1939.7 2000.9 C +1940.2 2001.9 1943 2001.8 1941.4 2003.3 C +1940.9 2003.6 1939.7 2003.6 1939.7 2004.3 C +1940.5 2003.9 1941.6 2003.7 1942.1 2003.1 C +f +S +n +vmrs +1967.1 1998.5 m +1967 1998.6 1966.8 1998.8 1966.6 1998.8 C +1966.7 1999.8 1967.8 2000.2 1968.5 2000.7 C +1969.4 2000.2 1968.2 1998.8 1967.1 1998.5 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1936.4 1997.6 m +1936.7 1995.6 1935.8 1996.4 1936.1 1997.8 C +1936.9 1997.9 1936.5 1999.2 1937.6 1999 C +1937 1998.5 1936.8 1998 1936.4 1997.6 C +f +S +n +1975.5 1996.6 m +1975.2 1996.7 1975.1 1996.5 1975 1996.4 C +1975 1996.2 1975 1996.1 1975 1995.9 C +1973.9 1996.5 1972 1995.5 1971.2 1996.8 C +1971.2 1998.6 1977 1999.9 1975.5 1996.6 C +[0 0 0 0] vc +f +S +n +1949.3 2097.4 m +1950.3 2096.9 1951.2 2096.4 1952.2 2096 C +1951.2 2096.4 1950.3 2096.9 1949.3 2097.4 C +[0.4 0.4 0 0] vc +f +S +n +1960.8 2091.6 m +1961.7 2091.2 1962.6 2090.9 1963.5 2090.4 C +1962.6 2090.9 1961.7 2091.2 1960.8 2091.6 C +f +S +n +1964.4 2090 m +1965.7 2089.2 1967 2088.5 1968.3 2087.8 C +1967 2088.5 1965.7 2089.2 1964.4 2090 C +f +S +n +1976 2083.7 m +1976.3 2082.3 1975.2 2079.1 1976.9 2079.4 C +1978.8 2080.7 1980.3 2082.9 1982.2 2084.2 C +1980.6 2083.1 1978.2 2080.2 1976 2078.9 C +1975.6 2081.2 1977 2084.9 1973.8 2085.4 C +1972.2 2086.1 1970.7 2087 1969 2087.6 C +1971.4 2086.5 1974.1 2085.6 1976 2083.7 C +f +S +n +1983.9 2084.2 m +1984.8 2083.7 1985.8 2083.2 1986.8 2082.8 C +1985.8 2083.2 1984.8 2083.7 1983.9 2084.2 C +f +S +n +1995.4 2078.4 m +1996.3 2078 1997.1 2077.7 1998 2077.2 C +1997.1 2077.7 1996.3 2078 1995.4 2078.4 C +f +S +n +1999 2076.8 m +2000.3 2076 2001.6 2075.3 2002.8 2074.6 C +2001.6 2075.3 2000.3 2076 1999 2076.8 C +f +S +n +vmrs +1929.6 2065.7 m +1930.1 2065.6 1929.8 2068.6 1929.9 2070 C +1929.8 2068.6 1930.1 2067 1929.6 2065.7 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1906.6 2049.4 m +1906.6 2046.7 1906.6 2043.9 1906.6 2041.2 C +1906.6 2043.9 1906.6 2046.7 1906.6 2049.4 C +f +S +n +2016 2047.5 m +2014.8 2048 2013.5 2048.3 2012.4 2049.4 C +2013.5 2048.3 2014.8 2048 2016 2047.5 C +f +S +n +2016.5 2047.2 m +2017.3 2046.9 2018.1 2046.6 2018.9 2046.3 C +2018.1 2046.6 2017.3 2046.9 2016.5 2047.2 C +f +S +n +1912.4 2028.5 m +1911.8 2032.4 1912.4 2037.2 1911.9 2041.2 C +1911.5 2037.2 1911.7 2032.9 1911.6 2028.8 C +1911.6 2033.5 1911.6 2038.9 1911.6 2042.9 C +1912.5 2042.2 1911.6 2043.9 1912.6 2043.6 C +1912.9 2039.3 1913.1 2033.3 1912.4 2028.5 C +[0.21 0.21 0 0] vc +f +S +n +1906.8 2040.8 m +1906.8 2039 1906.8 2037.2 1906.8 2035.5 C +1906.8 2037.2 1906.8 2039 1906.8 2040.8 C +[0.4 0.4 0 0] vc +f +S +n +1905.9 2035.2 m +1904.9 2036.4 1903.7 2037.2 1902.3 2037.4 C +1903.7 2037.2 1904.9 2036.4 1905.9 2035.2 C +f +S +n +1906.1 2031.2 m +1907 2031.1 1906.4 2028 1906.6 2030.7 C +1905.5 2032.1 1904 2032.8 1902.5 2033.6 C +1903.9 2033.2 1905 2032.1 1906.1 2031.2 C +f +S +n +1908.3 2018.7 m +1905.2 2018.6 1907.1 2023.2 1906.6 2025.4 C +1906.8 2023 1905.9 2019.5 1908.3 2018.7 C +f +S +n +1889.6 1998 m +1889 2001.9 1889.6 2006.7 1889.1 2010.8 C +1888.7 2006.7 1888.9 2002.4 1888.8 1998.3 C +1888.8 2003 1888.8 2008.4 1888.8 2012.4 C +1889.7 2011.7 1888.8 2013.4 1889.8 2013.2 C +1890.1 2008.8 1890.3 2002.8 1889.6 1998 C +[0.21 0.21 0 0] vc +f +S +n +vmrs +1999 2001.4 m +2001 2000.3 2003 1999.2 2005 1998 C +2003 1999.2 2001 2000.3 1999 2001.4 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1916.2 1986 m +1915.7 1989.9 1916.3 1994.7 1915.7 1998.8 C +1915.3 1994.7 1915.5 1990.4 1915.5 1986.3 C +1915.5 1991 1915.5 1996.4 1915.5 2000.4 C +1916.3 1999.7 1915.5 2001.4 1916.4 2001.2 C +1916.7 1996.8 1917 1990.8 1916.2 1986 C +[0.21 0.21 0 0] vc +f +S +n +1886.9 1989.6 m +1887.8 1989.2 1888.7 1988.9 1889.6 1988.4 C +1888.7 1988.9 1887.8 1989.2 1886.9 1989.6 C +[0.4 0.4 0 0] vc +f +S +n +1892.4 1986.8 m +1895.1 1985.1 1897.9 1983.6 1900.6 1982 C +1897.9 1983.6 1895.1 1985.1 1892.4 1986.8 C +f +S +n +1907.3 1979.3 m +1908.5 1978.9 1909.7 1978.5 1910.9 1978.1 C +1909.7 1978.5 1908.5 1978.9 1907.3 1979.3 C +f +S +n +1938.5 1966.6 m +1942.6 1970.1 1945.9 1976.4 1951.7 1979.1 C +1946.2 1976.1 1943.1 1970.9 1938.5 1966.6 C +f +S +n +1955.1 1978.6 m +1955.9 1978.2 1956.7 1977.8 1957.5 1977.4 C +1956.7 1977.8 1955.9 1978.2 1955.1 1978.6 C +f +S +n +1913.6 1977.6 m +1914.5 1977.2 1915.3 1976.9 1916.2 1976.4 C +1915.3 1976.9 1914.5 1977.2 1913.6 1977.6 C +f +S +n +1919.1 1974.8 m +1921.8 1973.1 1924.5 1971.6 1927.2 1970 C +1924.5 1971.6 1921.8 1973.1 1919.1 1974.8 C +f +S +n +1963.5 1974.5 m +1964.5 1974 1965.6 1973.4 1966.6 1972.8 C +1965.6 1973.4 1964.5 1974 1963.5 1974.5 C +f +S +n +vmrs +1967.8 1972.4 m +1970 1971.2 1972.1 1970 1974.3 1968.8 C +1972.1 1970 1970 1971.2 1967.8 1972.4 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1934 1967.3 m +1935.2 1966.9 1936.4 1966.5 1937.6 1966.1 C +1936.4 1966.5 1935.2 1966.9 1934 1967.3 C +f +S +n +1928.9 2061.2 m +1928.9 2059.2 1928.9 2057.3 1928.9 2055.4 C +1928.9 2057.3 1928.9 2059.2 1928.9 2061.2 C +[0.21 0.21 0 0] vc +f +S +n +1917.2 2047 m +1917.8 2046.5 1919.6 2046.8 1920 2047.2 C +1920 2046.5 1920.9 2046.8 1921 2046.3 C +1921.9 2047.3 1921.3 2044.1 1921.5 2044.1 C +1919.7 2044.8 1915.7 2043.5 1916.2 2046 C +1916.2 2048.3 1917 2045.9 1917.2 2047 C +[0 0 0 0] vc +f +S +n +1922 2044.1 m +1923.5 2043.2 1927 2045.4 1927.5 2042.9 C +1927.1 2042.6 1927.3 2040.9 1927.2 2041.5 C +1924.9 2042.3 1920.9 2040.6 1922 2044.1 C +f +S +n +1934.9 2043.9 m +1935.2 2043.4 1934.4 2042.7 1934 2042.2 C +1933.2 2041.8 1932.4 2042.8 1932.8 2043.2 C +1932.9 2044 1934.3 2043.3 1934.9 2043.9 C +f +S +n +1906.1 2030.7 m +1906.1 2028.8 1906.1 2027 1906.1 2025.2 C +1906.1 2027 1906.1 2028.8 1906.1 2030.7 C +[0.21 0.21 0 0] vc +f +S +n +1932.8 2018.7 m +1932.8 2016.8 1932.8 2014.8 1932.8 2012.9 C +1932.8 2014.8 1932.8 2016.8 1932.8 2018.7 C +f +S +n +1894.4 2016.5 m +1895 2016 1896.8 2016.3 1897.2 2016.8 C +1897.2 2016 1898.1 2016.3 1898.2 2015.8 C +1899.1 2016.8 1898.5 2013.6 1898.7 2013.6 C +1896.9 2014.4 1892.9 2013 1893.4 2015.6 C +1893.4 2017.8 1894.2 2015.4 1894.4 2016.5 C +[0 0 0 0] vc +f +S +n +1899.2 2013.6 m +1900.7 2012.7 1904.2 2014.9 1904.7 2012.4 C +1904.3 2012.1 1904.5 2010.5 1904.4 2011 C +1902.1 2011.8 1898.1 2010.1 1899.2 2013.6 C +f +S +n +vmrs +1912.1 2013.4 m +1912.4 2012.9 1911.6 2012.3 1911.2 2011.7 C +1910.4 2011.4 1909.6 2012.3 1910 2012.7 C +1910.1 2013.5 1911.5 2012.9 1912.1 2013.4 C +[0 0 0 0] vc +f +0.4 w +2 J +2 M +S +n +1921 2004.5 m +1921.6 2004 1923.4 2004.3 1923.9 2004.8 C +1923.8 2004 1924.8 2004.3 1924.8 2003.8 C +1925.7 2004.8 1925.1 2001.6 1925.3 2001.6 C +1923.6 2002.4 1919.6 2001 1920 2003.6 C +1920 2005.8 1920.8 2003.4 1921 2004.5 C +f +S +n +1925.8 2001.6 m +1927.3 2000.7 1930.8 2002.9 1931.3 2000.4 C +1930.9 2000.1 1931.1 1998.5 1931.1 1999 C +1928.7 1999.8 1924.8 1998.1 1925.8 2001.6 C +f +S +n +1938.8 2001.4 m +1939 2000.9 1938.2 2000.3 1937.8 1999.7 C +1937.1 1999.4 1936.2 2000.3 1936.6 2000.7 C +1936.7 2001.5 1938.1 2000.9 1938.8 2001.4 C +f +S +n +1908.6691 2008.1348 m +1897.82 2010.0477 L +1894.1735 1989.3671 L +1905.0226 1987.4542 L +1908.6691 2008.1348 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1895.041763 1994.291153 m +0 0 32 0 0 (l) ts +} +true +[0 0 0 1]sts +Q +1979.2185 1991.7809 m +1960.6353 1998.5452 L +1953.4532 1978.8124 L +1972.0363 1972.0481 L +1979.2185 1991.7809 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [18.793335 -6.84082 6.84021 18.793335 0 0] makesetfont +1955.163254 1983.510773 m +0 0 32 0 0 (\256) ts +} +true +[0 0 0 1]sts +Q +1952.1544 2066.5423 m +1938.0739 2069.025 L +1934.4274 2048.3444 L +1948.5079 2045.8617 L +1952.1544 2066.5423 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1935.29567 2053.268433 m +0 0 32 0 0 (") ts +} +true +[0 0 0 1]sts +Q +1931.7231 2043.621 m +1919.3084 2048.14 L +1910.6898 2024.4607 L +1923.1046 2019.9417 L +1931.7231 2043.621 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [22.552002 -8.208984 8.208252 22.552002 0 0] makesetfont +1912.741867 2030.098648 m +0 0 32 0 0 (=) ts +} +true +[0 0 0 1]sts +Q +1944 2024.5 m +1944 2014 L +0.8504 w +0 J +3.863693 M +[0 0 0 1] vc +false setoverprint +S +n +1944.25 2019.1673 m +1952.5 2015.9173 L +S +n +1931.0787 2124.423 m +1855.5505 2043.4285 L +1871.0419 2013.0337 L +1946.5701 2094.0282 L +1931.0787 2124.423 L +n +q +_bfh +%%IncludeResource: font ZapfHumanist601BT-Bold +_efh +{ +f1 [22.155762 23.759277 -14.753906 28.947754 0 0] makesetfont +1867.35347 2020.27063 m +0 0 32 0 0 (Isabelle) ts +} +true +[0 0 0 1]sts +Q +1933.5503 1996.9547 m +1922.7012 1998.8677 L +1919.0547 1978.1871 L +1929.9038 1976.2741 L +1933.5503 1996.9547 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1919.922913 1983.111069 m +0 0 32 0 0 (b) ts +} +true +[0 0 0 1]sts +Q +2006.3221 2025.7184 m +1993.8573 2027.9162 L +1990.2108 2007.2356 L +2002.6756 2005.0378 L +2006.3221 2025.7184 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1991.07901 2012.159653 m +0 0 32 0 0 (a) ts +} +true +[0 0 0 1]sts +Q +vmrs +2030.0624 2094.056 m +1956.3187 2120.904 L +1956.321 2095.3175 L +2030.0647 2068.4695 L +2030.0624 2094.056 L +n +q +_bfh +%%IncludeResource: font ZapfHumanist601BT-Bold +_efh +{ +f1 [22.898804 -8.336792 -0.002197 24.368408 0 0] makesetfont +1956.320496 2101.409561 m +0 0 32 0 0 (Isar) ts +} +true +[0 0 0 1]sts +Q +vmr +vmr +end +%%Trailer +%%DocumentNeededResources: font Symbol +%%+ font ZapfHumanist601BT-Bold +%%DocumentFonts: Symbol +%%+ ZapfHumanist601BT-Bold +%%DocumentNeededFonts: Symbol +%%+ ZapfHumanist601BT-Bold diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/isabelle_isar.pdf Binary file doc-src/Functions/isabelle_isar.pdf has changed diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/mathpartir.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/mathpartir.sty Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,421 @@ +% Mathpartir --- Math Paragraph for Typesetting Inference Rules +% +% Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy +% +% Author : Didier Remy +% Version : 1.2.0 +% Bug Reports : to author +% Web Site : http://pauillac.inria.fr/~remy/latex/ +% +% Mathpartir is free software; you can redistribute it and/or modify +% it under the terms of the GNU General Public License as published by +% the Free Software Foundation; either version 2, or (at your option) +% any later version. +% +% Mathpartir is distributed in the hope that it will be useful, +% but WITHOUT ANY WARRANTY; without even the implied warranty of +% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the +% GNU General Public License for more details +% (http://pauillac.inria.fr/~remy/license/GPL). +% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +% File mathpartir.sty (LaTeX macros) +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +\NeedsTeXFormat{LaTeX2e} +\ProvidesPackage{mathpartir} + [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules] + +%% + +%% Identification +%% Preliminary declarations + +\RequirePackage {keyval} + +%% Options +%% More declarations + +%% PART I: Typesetting maths in paragraphe mode + +\newdimen \mpr@tmpdim + +% To ensure hevea \hva compatibility, \hva should expands to nothing +% in mathpar or in inferrule +\let \mpr@hva \empty + +%% normal paragraph parametters, should rather be taken dynamically +\def \mpr@savepar {% + \edef \MathparNormalpar + {\noexpand \lineskiplimit \the\lineskiplimit + \noexpand \lineskip \the\lineskip}% + } + +\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em} +\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em} +\def \mpr@lineskip {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em} +\let \MathparLineskip \mpr@lineskip +\def \mpr@paroptions {\MathparLineskip} +\let \mpr@prebindings \relax + +\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em + +\def \mpr@goodbreakand + {\hskip -\mpr@andskip \penalty -1000\hskip \mpr@andskip} +\def \mpr@and {\hskip \mpr@andskip} +\def \mpr@andcr {\penalty 50\mpr@and} +\def \mpr@cr {\penalty -10000\mpr@and} +\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10} + +\def \mpr@bindings {% + \let \and \mpr@andcr + \let \par \mpr@andcr + \let \\\mpr@cr + \let \eqno \mpr@eqno + \let \hva \mpr@hva + } +\let \MathparBindings \mpr@bindings + +% \@ifundefined {ignorespacesafterend} +% {\def \ignorespacesafterend {\aftergroup \ignorespaces} + +\newenvironment{mathpar}[1][] + {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering + \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else + \noindent $\displaystyle\fi + \MathparBindings} + {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend} + +% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum +% \wd0 < \hsize $$\box0$$\else \bmathpar #1\emathpar \fi} + +%%% HOV BOXES + +\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip + \vbox \bgroup \tabskip 0em \let \\ \cr + \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup + \egroup} + +\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize + \box0\else \mathvbox {#1}\fi} + + +%% Part II -- operations on lists + +\newtoks \mpr@lista +\newtoks \mpr@listb + +\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter +{#2}\edef #2{\the \mpr@lista \the \mpr@listb}} + +\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter +{#2}\edef #2{\the \mpr@listb\the\mpr@lista}} + +\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb +\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}} + +\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2} +\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}} + +\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2} +\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}} + +\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}% + \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the + \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty + \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop + \mpr@flatten \mpr@all \mpr@to \mpr@one + \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof + \mpr@all \mpr@stripend + \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi + \ifx 1\mpr@isempty + \repeat +} + +\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty + \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp} + +%% Part III -- Type inference rules + +\newif \if@premisse +\newbox \mpr@hlist +\newbox \mpr@vlist +\newif \ifmpr@center \mpr@centertrue +\def \mpr@htovlist {% + \setbox \mpr@hlist + \hbox {\strut + \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi + \unhbox \mpr@hlist}% + \setbox \mpr@vlist + \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist + \else \unvbox \mpr@vlist \box \mpr@hlist + \fi}% +} +% OLD version +% \def \mpr@htovlist {% +% \setbox \mpr@hlist +% \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}% +% \setbox \mpr@vlist +% \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist +% \else \unvbox \mpr@vlist \box \mpr@hlist +% \fi}% +% } + +\def \mpr@item #1{$\displaystyle #1$} +\def \mpr@sep{2em} +\def \mpr@blank { } +\def \mpr@hovbox #1#2{\hbox + \bgroup + \ifx #1T\@premissetrue + \else \ifx #1B\@premissefalse + \else + \PackageError{mathpartir} + {Premisse orientation should either be T or B} + {Fatal error in Package}% + \fi \fi + \def \@test {#2}\ifx \@test \mpr@blank\else + \setbox \mpr@hlist \hbox {}% + \setbox \mpr@vlist \vbox {}% + \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi + \let \@hvlist \empty \let \@rev \empty + \mpr@tmpdim 0em + \expandafter \mpr@makelist #2\mpr@to \mpr@flat + \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi + \def \\##1{% + \def \@test {##1}\ifx \@test \empty + \mpr@htovlist + \mpr@tmpdim 0em %%% last bug fix not extensively checked + \else + \setbox0 \hbox{\mpr@item {##1}}\relax + \advance \mpr@tmpdim by \wd0 + %\mpr@tmpdim 1.02\mpr@tmpdim + \ifnum \mpr@tmpdim < \hsize + \ifnum \wd\mpr@hlist > 0 + \if@premisse + \setbox \mpr@hlist + \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}% + \else + \setbox \mpr@hlist + \hbox {\unhbox \mpr@hlist \hskip \mpr@sep \unhbox0}% + \fi + \else + \setbox \mpr@hlist \hbox {\unhbox0}% + \fi + \else + \ifnum \wd \mpr@hlist > 0 + \mpr@htovlist + \mpr@tmpdim \wd0 + \fi + \setbox \mpr@hlist \hbox {\unhbox0}% + \fi + \advance \mpr@tmpdim by \mpr@sep + \fi + }% + \@rev + \mpr@htovlist + \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist + \fi + \egroup +} + +%%% INFERENCE RULES + +\@ifundefined{@@over}{% + \let\@@over\over % fallback if amsmath is not loaded + \let\@@overwithdelims\overwithdelims + \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims + \let\@@above\above \let\@@abovewithdelims\abovewithdelims + }{} + +%% The default + +\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em + $\displaystyle {#1\mpr@over #2}$}} +\let \mpr@fraction \mpr@@fraction + +%% A generic solution to arrow + +\def \mpr@make@fraction #1#2#3#4#5{\hbox {% + \def \mpr@tail{#1}% + \def \mpr@body{#2}% + \def \mpr@head{#3}% + \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}% + \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}% + \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}% + \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax + \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax + \setbox0=\hbox {$\box1 \@@atop \box2$}% + \dimen0=\wd0\box0 + \box0 \hskip -\dimen0\relax + \hbox to \dimen0 {$% + \mathrel{\mpr@tail}\joinrel + \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}% + $}}} + +%% Old stuff should be removed in next version +\def \mpr@@reduce #1#2{\hbox + {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}} +\def \mpr@@rewrite #1#2#3{\hbox + {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}} +\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}} + +\def \mpr@empty {} +\def \mpr@inferrule + {\bgroup + \ifnum \linewidth<\hsize \hsize \linewidth\fi + \mpr@rulelineskip + \let \and \qquad + \let \hva \mpr@hva + \let \@rulename \mpr@empty + \let \@rule@options \mpr@empty + \let \mpr@over \@@over + \mpr@inferrule@} +\newcommand {\mpr@inferrule@}[3][] + {\everymath={\displaystyle}% + \def \@test {#2}\ifx \empty \@test + \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}% + \else + \def \@test {#3}\ifx \empty \@test + \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}% + \else + \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}% + \fi \fi + \def \@test {#1}\ifx \@test\empty \box0 + \else \vbox +%%% Suggestion de Francois pour les etiquettes longues +%%% {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi + {\hbox {\RefTirName {#1}}\box0}\fi + \egroup} + +\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}} + +% They are two forms +% \inferrule [label]{[premisses}{conclusions} +% or +% \inferrule* [options]{[premisses}{conclusions} +% +% Premisses and conclusions are lists of elements separated by \\ +% Each \\ produces a break, attempting horizontal breaks if possible, +% and vertical breaks if needed. +% +% An empty element obtained by \\\\ produces a vertical break in all cases. +% +% The former rule is aligned on the fraction bar. +% The optional label appears on top of the rule +% The second form to be used in a derivation tree is aligned on the last +% line of its conclusion +% +% The second form can be parameterized, using the key=val interface. The +% folloiwng keys are recognized: +% +% width set the width of the rule to val +% narrower set the width of the rule to val\hsize +% before execute val at the beginning/left +% lab put a label [Val] on top of the rule +% lskip add negative skip on the right +% left put a left label [Val] +% Left put a left label [Val], ignoring its width +% right put a right label [Val] +% Right put a right label [Val], ignoring its width +% leftskip skip negative space on the left-hand side +% rightskip skip negative space on the right-hand side +% vdots lift the rule by val and fill vertical space with dots +% after execute val at the end/right +% +% Note that most options must come in this order to avoid strange +% typesetting (in particular leftskip must preceed left and Left and +% rightskip must follow Right or right; vdots must come last +% or be only followed by rightskip. +% + +%% Keys that make sence in all kinds of rules +\def \mprset #1{\setkeys{mprset}{#1}} +\define@key {mprset}{flushleft}[]{\mpr@centerfalse} +\define@key {mprset}{center}[]{\mpr@centertrue} +\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} +\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1} +\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} + +\newbox \mpr@right +\define@key {mpr}{flushleft}[]{\mpr@centerfalse} +\define@key {mpr}{center}[]{\mpr@centertrue} +\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} +\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1} +\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} +\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax + \advance \hsize by -\wd0\box0} +\define@key {mpr}{width}{\hsize #1} +\define@key {mpr}{sep}{\def\mpr@sep{#1}} +\define@key {mpr}{before}{#1} +\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} +\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} +\define@key {mpr}{narrower}{\hsize #1\hsize} +\define@key {mpr}{leftskip}{\hskip -#1} +\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce} +\define@key {mpr}{rightskip} + {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}} +\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax + \advance \hsize by -\wd0\box0} +\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax + \advance \hsize by -\wd0\box0} +\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}} +\define@key {mpr}{right} + {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0 + \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} +\define@key {mpr}{RIGHT} + {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0 + \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} +\define@key {mpr}{Right} + {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}} +\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}} +\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}} + +\newdimen \rule@dimen +\newcommand \mpr@inferstar@ [3][]{\setbox0 + \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax + \setbox \mpr@right \hbox{}% + $\setkeys{mpr}{#1}% + \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else + \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi + \box \mpr@right \mpr@vdots$} + \setbox1 \hbox {\strut} + \rule@dimen \dp0 \advance \rule@dimen by -\dp1 + \raise \rule@dimen \box0} + +\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}} +\newcommand \mpr@err@skipargs[3][]{} +\def \mpr@inferstar*{\ifmmode + \let \@do \mpr@inferstar@ + \else + \let \@do \mpr@err@skipargs + \PackageError {mathpartir} + {\string\inferrule* can only be used in math mode}{}% + \fi \@do} + + +%%% Exports + +% Envirnonment mathpar + +\let \inferrule \mpr@infer + +% make a short name \infer is not already defined +\@ifundefined {infer}{\let \infer \mpr@infer}{} + +\def \TirNameStyle #1{\small \textsc{#1}} +\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}} +\let \TirName \tir@name +\let \DefTirName \TirName +\let \RefTirName \TirName + +%%% Other Exports + +% \let \listcons \mpr@cons +% \let \listsnoc \mpr@snoc +% \let \listhead \mpr@head +% \let \listmake \mpr@makelist + + + + +\endinput diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/Functions/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/style.sty Tue Mar 03 11:00:51 2009 +0100 @@ -0,0 +1,46 @@ +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} +\newcommand{\chref}[1]{chapter~\ref{#1}} +\newcommand{\figref}[1]{figure~\ref{#1}} + +%% math +\newcommand{\text}[1]{\mbox{#1}} +\newcommand{\isasymvartheta}{\isamath{\theta}} +\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}} + +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +\pagestyle{headings} +\sloppy +\binperiod +\underscoreon + +\renewcommand{\isadigit}[1]{\isamath{#1}} + +\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} + +\isafoldtag{FIXME} +\isakeeptag{mlref} +\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} +\renewcommand{\endisatagmlref}{\endgroup} + +\newcommand{\isasymGUESS}{\isakeyword{guess}} +\newcommand{\isasymOBTAIN}{\isakeyword{obtain}} +\newcommand{\isasymTHEORY}{\isakeyword{theory}} +\newcommand{\isasymUSES}{\isakeyword{uses}} +\newcommand{\isasymEND}{\isakeyword{end}} +\newcommand{\isasymCONSTS}{\isakeyword{consts}} +\newcommand{\isasymDEFS}{\isakeyword{defs}} +\newcommand{\isasymTHEOREM}{\isakeyword{theorem}} +\newcommand{\isasymDEFINITION}{\isakeyword{definition}} + +\isabellestyle{it} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "implementation" +%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/IsaMakefile --- a/doc-src/IsarAdvanced/Classes/IsaMakefile Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ - -## targets - -default: Thy -images: -test: Thy - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log - -USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document - - -## Thy - -THY = $(LOG)/HOL-Thy.gz - -Thy: $(THY) - -$(THY): Thy/ROOT.ML Thy/Setup.thy Thy/Classes.thy - @$(USEDIR) HOL Thy - - -## clean - -clean: - @rm -f $(THY) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/Makefile --- a/doc-src/IsarAdvanced/Classes/Makefile Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = classes - -FILES = $(NAME).tex classes.tex Thy/document/Classes.tex \ - style.sty ../../iman.sty ../../extra.sty ../../isar.sty \ - ../../isabelle.sty ../../isabellesym.sty ../../pdfsetup.sty \ - ../../manual.bib ../../proof.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf - $(PDFLATEX) $(NAME) - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/Thy/Classes.thy --- a/doc-src/IsarAdvanced/Classes/Thy/Classes.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,635 +0,0 @@ -theory Classes -imports Main Setup -begin - -chapter {* Haskell-style classes with Isabelle/Isar *} - -section {* Introduction *} - -text {* - Type classes were introduces by Wadler and Blott \cite{wadler89how} - into the Haskell language, to allow for a reasonable implementation - of overloading\footnote{throughout this tutorial, we are referring - to classical Haskell 1.0 type classes, not considering - later additions in expressiveness}. - As a canonical example, a polymorphic equality function - @{text "eq \ \ \ \ \ bool"} which is overloaded on different - types for @{text "\"}, which is achieved by splitting introduction - of the @{text eq} function from its overloaded definitions by means - of @{text class} and @{text instance} declarations: - - \begin{quote} - - \noindent@{text "class eq where"}\footnote{syntax here is a kind of isabellized Haskell} \\ - \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} - - \medskip\noindent@{text "instance nat \ eq where"} \\ - \hspace*{2ex}@{text "eq 0 0 = True"} \\ - \hspace*{2ex}@{text "eq 0 _ = False"} \\ - \hspace*{2ex}@{text "eq _ 0 = False"} \\ - \hspace*{2ex}@{text "eq (Suc n) (Suc m) = eq n m"} - - \medskip\noindent@{text "instance (\\eq, \\eq) pair \ eq where"} \\ - \hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2"} - - \medskip\noindent@{text "class ord extends eq where"} \\ - \hspace*{2ex}@{text "less_eq \ \ \ \ \ bool"} \\ - \hspace*{2ex}@{text "less \ \ \ \ \ bool"} - - \end{quote} - - \noindent Type variables are annotated with (finitely many) classes; - these annotations are assertions that a particular polymorphic type - provides definitions for overloaded functions. - - Indeed, type classes not only allow for simple overloading - but form a generic calculus, an instance of order-sorted - algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. - - From a software engeneering point of view, type classes - roughly correspond to interfaces in object-oriented languages like Java; - so, it is naturally desirable that type classes do not only - provide functions (class parameters) but also state specifications - implementations must obey. For example, the @{text "class eq"} - above could be given the following specification, demanding that - @{text "class eq"} is an equivalence relation obeying reflexivity, - symmetry and transitivity: - - \begin{quote} - - \noindent@{text "class eq where"} \\ - \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \\ - @{text "satisfying"} \\ - \hspace*{2ex}@{text "refl: eq x x"} \\ - \hspace*{2ex}@{text "sym: eq x y \ eq x y"} \\ - \hspace*{2ex}@{text "trans: eq x y \ eq y z \ eq x z"} - - \end{quote} - - \noindent From a theoretic point of view, type classes are lightweight - modules; Haskell type classes may be emulated by - SML functors \cite{classes_modules}. - Isabelle/Isar offers a discipline of type classes which brings - all those aspects together: - - \begin{enumerate} - \item specifying abstract parameters together with - corresponding specifications, - \item instantiating those abstract parameters by a particular - type - \item in connection with a ``less ad-hoc'' approach to overloading, - \item with a direct link to the Isabelle module system - (aka locales \cite{kammueller-locales}). - \end{enumerate} - - \noindent Isar type classes also directly support code generation - in a Haskell like fashion. - - This tutorial demonstrates common elements of structured specifications - and abstract reasoning with type classes by the algebraic hierarchy of - semigroups, monoids and groups. Our background theory is that of - Isabelle/HOL \cite{isa-tutorial}, for which some - familiarity is assumed. - - Here we merely present the look-and-feel for end users. - Internally, those are mapped to more primitive Isabelle concepts. - See \cite{Haftmann-Wenzel:2006:classes} for more detail. -*} - -section {* A simple algebra example \label{sec:example} *} - -subsection {* Class definition *} - -text {* - Depending on an arbitrary type @{text "\"}, class @{text - "semigroup"} introduces a binary operator @{text "(\)"} that is - assumed to be associative: -*} - -class %quote semigroup = - fixes mult :: "\ \ \ \ \" (infixl "\" 70) - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - -text {* - \noindent This @{command class} specification consists of two - parts: the \qn{operational} part names the class parameter - (@{element "fixes"}), the \qn{logical} part specifies properties on them - (@{element "assumes"}). The local @{element "fixes"} and - @{element "assumes"} are lifted to the theory toplevel, - yielding the global - parameter @{term [source] "mult \ \\semigroup \ \ \ \"} and the - global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y - z \ \\semigroup. (x \ y) \ z = x \ (y \ z)"}. -*} - - -subsection {* Class instantiation \label{sec:class_inst} *} - -text {* - The concrete type @{typ int} is made a @{class semigroup} - instance by providing a suitable definition for the class parameter - @{text "(\)"} and a proof for the specification of @{fact assoc}. - This is accomplished by the @{command instantiation} target: -*} - -instantiation %quote int :: semigroup -begin - -definition %quote - mult_int_def: "i \ j = i + (j\int)" - -instance %quote proof - fix i j k :: int have "(i + j) + k = i + (j + k)" by simp - then show "(i \ j) \ k = i \ (j \ k)" - unfolding mult_int_def . -qed - -end %quote - -text {* - \noindent @{command instantiation} allows to define class parameters - at a particular instance using common specification tools (here, - @{command definition}). The concluding @{command instance} - opens a proof that the given parameters actually conform - to the class specification. Note that the first proof step - is the @{method default} method, - which for such instance proofs maps to the @{method intro_classes} method. - This boils down an instance judgement to the relevant primitive - proof goals and should conveniently always be the first method applied - in an instantiation proof. - - From now on, the type-checker will consider @{typ int} - as a @{class semigroup} automatically, i.e.\ any general results - are immediately available on concrete instances. - - \medskip Another instance of @{class semigroup} are the natural numbers: -*} - -instantiation %quote nat :: semigroup -begin - -primrec %quote mult_nat where - "(0\nat) \ n = n" - | "Suc m \ n = Suc (m \ n)" - -instance %quote proof - fix m n q :: nat - show "m \ n \ q = m \ (n \ q)" - by (induct m) auto -qed - -end %quote - -text {* - \noindent Note the occurence of the name @{text mult_nat} - in the primrec declaration; by default, the local name of - a class operation @{text f} to instantiate on type constructor - @{text \} are mangled as @{text f_\}. In case of uncertainty, - these names may be inspected using the @{command "print_context"} command - or the corresponding ProofGeneral button. -*} - -subsection {* Lifting and parametric types *} - -text {* - Overloaded definitions giving on class instantiation - may include recursion over the syntactic structure of types. - As a canonical example, we model product semigroups - using our simple algebra: -*} - -instantiation %quote * :: (semigroup, semigroup) semigroup -begin - -definition %quote - mult_prod_def: "p\<^isub>1 \ p\<^isub>2 = (fst p\<^isub>1 \ fst p\<^isub>2, snd p\<^isub>1 \ snd p\<^isub>2)" - -instance %quote proof - fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\\semigroup \ \\semigroup" - show "p\<^isub>1 \ p\<^isub>2 \ p\<^isub>3 = p\<^isub>1 \ (p\<^isub>2 \ p\<^isub>3)" - unfolding mult_prod_def by (simp add: assoc) -qed - -end %quote - -text {* - \noindent Associativity from product semigroups is - established using - the definition of @{text "(\)"} on products and the hypothetical - associativity of the type components; these hypotheses - are facts due to the @{class semigroup} constraints imposed - on the type components by the @{command instance} proposition. - Indeed, this pattern often occurs with parametric types - and type classes. -*} - - -subsection {* Subclassing *} - -text {* - We define a subclass @{text monoidl} (a semigroup with a left-hand neutral) - by extending @{class semigroup} - with one additional parameter @{text neutral} together - with its property: -*} - -class %quote monoidl = semigroup + - fixes neutral :: "\" ("\") - assumes neutl: "\ \ x = x" - -text {* - \noindent Again, we prove some instances, by - providing suitable parameter definitions and proofs for the - additional specifications. Observe that instantiations - for types with the same arity may be simultaneous: -*} - -instantiation %quote nat and int :: monoidl -begin - -definition %quote - neutral_nat_def: "\ = (0\nat)" - -definition %quote - neutral_int_def: "\ = (0\int)" - -instance %quote proof - fix n :: nat - show "\ \ n = n" - unfolding neutral_nat_def by simp -next - fix k :: int - show "\ \ k = k" - unfolding neutral_int_def mult_int_def by simp -qed - -end %quote - -instantiation %quote * :: (monoidl, monoidl) monoidl -begin - -definition %quote - neutral_prod_def: "\ = (\, \)" - -instance %quote proof - fix p :: "\\monoidl \ \\monoidl" - show "\ \ p = p" - unfolding neutral_prod_def mult_prod_def by (simp add: neutl) -qed - -end %quote - -text {* - \noindent Fully-fledged monoids are modelled by another subclass - which does not add new parameters but tightens the specification: -*} - -class %quote monoid = monoidl + - assumes neutr: "x \ \ = x" - -instantiation %quote nat and int :: monoid -begin - -instance %quote proof - fix n :: nat - show "n \ \ = n" - unfolding neutral_nat_def by (induct n) simp_all -next - fix k :: int - show "k \ \ = k" - unfolding neutral_int_def mult_int_def by simp -qed - -end %quote - -instantiation %quote * :: (monoid, monoid) monoid -begin - -instance %quote proof - fix p :: "\\monoid \ \\monoid" - show "p \ \ = p" - unfolding neutral_prod_def mult_prod_def by (simp add: neutr) -qed - -end %quote - -text {* - \noindent To finish our small algebra example, we add a @{text group} class - with a corresponding instance: -*} - -class %quote group = monoidl + - fixes inverse :: "\ \ \" ("(_\
)" [1000] 999) - assumes invl: "x\
\ x = \" - -instantiation %quote int :: group -begin - -definition %quote - inverse_int_def: "i\
= - (i\int)" - -instance %quote proof - fix i :: int - have "-i + i = 0" by simp - then show "i\
\ i = \" - unfolding mult_int_def neutral_int_def inverse_int_def . -qed - -end %quote - - -section {* Type classes as locales *} - -subsection {* A look behind the scene *} - -text {* - The example above gives an impression how Isar type classes work - in practice. As stated in the introduction, classes also provide - a link to Isar's locale system. Indeed, the logical core of a class - is nothing else than a locale: -*} - -class %quote idem = - fixes f :: "\ \ \" - assumes idem: "f (f x) = f x" - -text {* - \noindent essentially introduces the locale -*} setup %invisible {* Sign.add_path "foo" *} - -locale %quote idem = - fixes f :: "\ \ \" - assumes idem: "f (f x) = f x" - -text {* \noindent together with corresponding constant(s): *} - -consts %quote f :: "\ \ \" - -text {* - \noindent The connection to the type system is done by means - of a primitive axclass -*} setup %invisible {* Sign.add_path "foo" *} - -axclass %quote idem < type - idem: "f (f x) = f x" setup %invisible {* Sign.parent_path *} - -text {* \noindent together with a corresponding interpretation: *} - -interpretation %quote idem_class: - idem "f \ (\\idem) \ \" -proof qed (rule idem) - -text {* - \noindent This gives you at hand the full power of the Isabelle module system; - conclusions in locale @{text idem} are implicitly propagated - to class @{text idem}. -*} setup %invisible {* Sign.parent_path *} - -subsection {* Abstract reasoning *} - -text {* - Isabelle locales enable reasoning at a general level, while results - are implicitly transferred to all instances. For example, we can - now establish the @{text "left_cancel"} lemma for groups, which - states that the function @{text "(x \)"} is injective: -*} - -lemma %quote (in group) left_cancel: "x \ y = x \ z \ y = z" -proof - assume "x \ y = x \ z" - then have "x\
\ (x \ y) = x\
\ (x \ z)" by simp - then have "(x\
\ x) \ y = (x\
\ x) \ z" using assoc by simp - then show "y = z" using neutl and invl by simp -next - assume "y = z" - then show "x \ y = x \ z" by simp -qed - -text {* - \noindent Here the \qt{@{keyword "in"} @{class group}} target specification - indicates that the result is recorded within that context for later - use. This local theorem is also lifted to the global one @{fact - "group.left_cancel:"} @{prop [source] "\x y z \ \\group. x \ y = x \ - z \ y = z"}. Since type @{text "int"} has been made an instance of - @{text "group"} before, we may refer to that fact as well: @{prop - [source] "\x y z \ int. x \ y = x \ z \ y = z"}. -*} - - -subsection {* Derived definitions *} - -text {* - Isabelle locales support a concept of local definitions - in locales: -*} - -primrec %quote (in monoid) pow_nat :: "nat \ \ \ \" where - "pow_nat 0 x = \" - | "pow_nat (Suc n) x = x \ pow_nat n x" - -text {* - \noindent If the locale @{text group} is also a class, this local - definition is propagated onto a global definition of - @{term [source] "pow_nat \ nat \ \\monoid \ \\monoid"} - with corresponding theorems - - @{thm pow_nat.simps [no_vars]}. - - \noindent As you can see from this example, for local - definitions you may use any specification tool - which works together with locales (e.g. \cite{krauss2006}). -*} - - -subsection {* A functor analogy *} - -text {* - We introduced Isar classes by analogy to type classes - functional programming; if we reconsider this in the - context of what has been said about type classes and locales, - we can drive this analogy further by stating that type - classes essentially correspond to functors which have - a canonical interpretation as type classes. - Anyway, there is also the possibility of other interpretations. - For example, also @{text list}s form a monoid with - @{text append} and @{term "[]"} as operations, but it - seems inappropriate to apply to lists - the same operations as for genuinely algebraic types. - In such a case, we simply can do a particular interpretation - of monoids for lists: -*} - -interpretation %quote list_monoid!: monoid append "[]" - proof qed auto - -text {* - \noindent This enables us to apply facts on monoids - to lists, e.g. @{thm list_monoid.neutl [no_vars]}. - - When using this interpretation pattern, it may also - be appropriate to map derived definitions accordingly: -*} - -primrec %quote replicate :: "nat \ \ list \ \ list" where - "replicate 0 _ = []" - | "replicate (Suc n) xs = xs @ replicate n xs" - -interpretation %quote list_monoid!: monoid append "[]" where - "monoid.pow_nat append [] = replicate" -proof - - interpret monoid append "[]" .. - show "monoid.pow_nat append [] = replicate" - proof - fix n - show "monoid.pow_nat append [] n = replicate n" - by (induct n) auto - qed -qed intro_locales - - -subsection {* Additional subclass relations *} - -text {* - Any @{text "group"} is also a @{text "monoid"}; this - can be made explicit by claiming an additional - subclass relation, - together with a proof of the logical difference: -*} - -subclass %quote (in group) monoid -proof - fix x - from invl have "x\
\ x = \" by simp - with assoc [symmetric] neutl invl have "x\
\ (x \ \) = x\
\ x" by simp - with left_cancel show "x \ \ = x" by simp -qed - -text {* - \noindent The logical proof is carried out on the locale level. - Afterwards it is propagated - to the type system, making @{text group} an instance of - @{text monoid} by adding an additional edge - to the graph of subclass relations - (cf.\ \figref{fig:subclass}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){@{text semigroup}}} - \put(20,40){\makebox(0,0){@{text monoidl}}} - \put(00,20){\makebox(0,0){@{text monoid}}} - \put(40,00){\makebox(0,0){@{text group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(25,35){\vector(1,-3){10}} - \end{picture} - \hspace{8em} - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){@{text semigroup}}} - \put(20,40){\makebox(0,0){@{text monoidl}}} - \put(00,20){\makebox(0,0){@{text monoid}}} - \put(40,00){\makebox(0,0){@{text group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(05,15){\vector(3,-1){30}} - \end{picture} - \caption{Subclass relationship of monoids and groups: - before and after establishing the relationship - @{text "group \ monoid"}; transitive edges are left out.} - \label{fig:subclass} - \end{center} - \end{figure} -7 - For illustration, a derived definition - in @{text group} which uses @{text pow_nat}: -*} - -definition %quote (in group) pow_int :: "int \ \ \ \" where - "pow_int k x = (if k >= 0 - then pow_nat (nat k) x - else (pow_nat (nat (- k)) x)\
)" - -text {* - \noindent yields the global definition of - @{term [source] "pow_int \ int \ \\group \ \\group"} - with the corresponding theorem @{thm pow_int_def [no_vars]}. -*} - -subsection {* A note on syntax *} - -text {* - As a commodity, class context syntax allows to refer - to local class operations and their global counterparts - uniformly; type inference resolves ambiguities. For example: -*} - -context %quote semigroup -begin - -term %quote "x \ y" -- {* example 1 *} -term %quote "(x\nat) \ y" -- {* example 2 *} - -end %quote - -term %quote "x \ y" -- {* example 3 *} - -text {* - \noindent Here in example 1, the term refers to the local class operation - @{text "mult [\]"}, whereas in example 2 the type constraint - enforces the global class operation @{text "mult [nat]"}. - In the global context in example 3, the reference is - to the polymorphic global class operation @{text "mult [?\ \ semigroup]"}. -*} - -section {* Further issues *} - -subsection {* Type classes and code generation *} - -text {* - Turning back to the first motivation for type classes, - namely overloading, it is obvious that overloading - stemming from @{command class} statements and - @{command instantiation} - targets naturally maps to Haskell type classes. - The code generator framework \cite{isabelle-codegen} - takes this into account. Concerning target languages - lacking type classes (e.g.~SML), type classes - are implemented by explicit dictionary construction. - As example, let's go back to the power function: -*} - -definition %quote example :: int where - "example = pow_int 10 (-2)" - -text {* - \noindent This maps to Haskell as: -*} - -text %quote {*@{code_stmts example (Haskell)}*} - -text {* - \noindent The whole code in SML with explicit dictionary passing: -*} - -text %quote {*@{code_stmts example (SML)}*} - -subsection {* Inspecting the type class universe *} - -text {* - To facilitate orientation in complex subclass structures, - two diagnostics commands are provided: - - \begin{description} - - \item[@{command "print_classes"}] print a list of all classes - together with associated operations etc. - - \item[@{command "class_deps"}] visualizes the subclass relation - between all classes as a Hasse diagram. - - \end{description} -*} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/Thy/ROOT.ML --- a/doc-src/IsarAdvanced/Classes/Thy/ROOT.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ - -(* $Id$ *) - -no_document use_thy "Setup"; - -use_thy "Classes"; diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/Thy/Setup.thy --- a/doc-src/IsarAdvanced/Classes/Thy/Setup.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,34 +0,0 @@ -theory Setup -imports Main Code_Integer -uses - "../../../antiquote_setup" - "../../../more_antiquote" -begin - -ML {* Code_Target.code_width := 74 *} - -syntax - "_alpha" :: "type" ("\") - "_alpha_ofsort" :: "sort \ type" ("\()\_" [0] 1000) - "_beta" :: "type" ("\") - "_beta_ofsort" :: "sort \ type" ("\()\_" [0] 1000) - -parse_ast_translation {* - let - fun alpha_ast_tr [] = Syntax.Variable "'a" - | alpha_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); - fun alpha_ofsort_ast_tr [ast] = - Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'a", ast] - | alpha_ofsort_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); - fun beta_ast_tr [] = Syntax.Variable "'b" - | beta_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); - fun beta_ofsort_ast_tr [ast] = - Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'b", ast] - | beta_ofsort_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); - in [ - ("_alpha", alpha_ast_tr), ("_alpha_ofsort", alpha_ofsort_ast_tr), - ("_beta", beta_ast_tr), ("_beta_ofsort", beta_ofsort_ast_tr) - ] end -*} - -end \ No newline at end of file diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex --- a/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1346 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Classes}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Classes\isanewline -\isakeyword{imports}\ Main\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Haskell-style classes with Isabelle/Isar% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Type classes were introduces by Wadler and Blott \cite{wadler89how} - into the Haskell language, to allow for a reasonable implementation - of overloading\footnote{throughout this tutorial, we are referring - to classical Haskell 1.0 type classes, not considering - later additions in expressiveness}. - As a canonical example, a polymorphic equality function - \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different - types for \isa{{\isasymalpha}}, which is achieved by splitting introduction - of the \isa{eq} function from its overloaded definitions by means - of \isa{class} and \isa{instance} declarations: - - \begin{quote} - - \noindent\isa{class\ eq\ where}\footnote{syntax here is a kind of isabellized Haskell} \\ - \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} - - \medskip\noindent\isa{instance\ nat\ {\isasymColon}\ eq\ where} \\ - \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isadigit{0}}\ {\isacharequal}\ True} \\ - \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ False} \\ - \hspace*{2ex}\isa{eq\ {\isacharunderscore}\ {\isadigit{0}}\ {\isacharequal}\ False} \\ - \hspace*{2ex}\isa{eq\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ eq\ n\ m} - - \medskip\noindent\isa{instance\ {\isacharparenleft}{\isasymalpha}{\isasymColon}eq{\isacharcomma}\ {\isasymbeta}{\isasymColon}eq{\isacharparenright}\ pair\ {\isasymColon}\ eq\ where} \\ - \hspace*{2ex}\isa{eq\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ eq\ x{\isadigit{1}}\ x{\isadigit{2}}\ {\isasymand}\ eq\ y{\isadigit{1}}\ y{\isadigit{2}}} - - \medskip\noindent\isa{class\ ord\ extends\ eq\ where} \\ - \hspace*{2ex}\isa{less{\isacharunderscore}eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ - \hspace*{2ex}\isa{less\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} - - \end{quote} - - \noindent Type variables are annotated with (finitely many) classes; - these annotations are assertions that a particular polymorphic type - provides definitions for overloaded functions. - - Indeed, type classes not only allow for simple overloading - but form a generic calculus, an instance of order-sorted - algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. - - From a software engeneering point of view, type classes - roughly correspond to interfaces in object-oriented languages like Java; - so, it is naturally desirable that type classes do not only - provide functions (class parameters) but also state specifications - implementations must obey. For example, the \isa{class\ eq} - above could be given the following specification, demanding that - \isa{class\ eq} is an equivalence relation obeying reflexivity, - symmetry and transitivity: - - \begin{quote} - - \noindent\isa{class\ eq\ where} \\ - \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ - \isa{satisfying} \\ - \hspace*{2ex}\isa{refl{\isacharcolon}\ eq\ x\ x} \\ - \hspace*{2ex}\isa{sym{\isacharcolon}\ eq\ x\ y\ {\isasymlongleftrightarrow}\ eq\ x\ y} \\ - \hspace*{2ex}\isa{trans{\isacharcolon}\ eq\ x\ y\ {\isasymand}\ eq\ y\ z\ {\isasymlongrightarrow}\ eq\ x\ z} - - \end{quote} - - \noindent From a theoretic point of view, type classes are lightweight - modules; Haskell type classes may be emulated by - SML functors \cite{classes_modules}. - Isabelle/Isar offers a discipline of type classes which brings - all those aspects together: - - \begin{enumerate} - \item specifying abstract parameters together with - corresponding specifications, - \item instantiating those abstract parameters by a particular - type - \item in connection with a ``less ad-hoc'' approach to overloading, - \item with a direct link to the Isabelle module system - (aka locales \cite{kammueller-locales}). - \end{enumerate} - - \noindent Isar type classes also directly support code generation - in a Haskell like fashion. - - This tutorial demonstrates common elements of structured specifications - and abstract reasoning with type classes by the algebraic hierarchy of - semigroups, monoids and groups. Our background theory is that of - Isabelle/HOL \cite{isa-tutorial}, for which some - familiarity is assumed. - - Here we merely present the look-and-feel for end users. - Internally, those are mapped to more primitive Isabelle concepts. - See \cite{Haftmann-Wenzel:2006:classes} for more detail.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{A simple algebra example \label{sec:example}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Class definition% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operator \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} that is - assumed to be associative:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ semigroup\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two - parts: the \qn{operational} part names the class parameter - (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them - (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and - \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, - yielding the global - parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the - global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Class instantiation \label{sec:class_inst}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The concrete type \isa{int} is made a \isa{semigroup} - instance by providing a suitable definition for the class parameter - \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}. - This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymotimes}\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} allows to define class parameters - at a particular instance using common specification tools (here, - \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} - opens a proof that the given parameters actually conform - to the class specification. Note that the first proof step - is the \hyperlink{method.default}{\mbox{\isa{default}}} method, - which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method. - This boils down an instance judgement to the relevant primitive - proof goals and should conveniently always be the first method applied - in an instantiation proof. - - From now on, the type-checker will consider \isa{int} - as a \isa{semigroup} automatically, i.e.\ any general results - are immediately available on concrete instances. - - \medskip Another instance of \isa{semigroup} are the natural numbers:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}m\ {\isasymotimes}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ m{\isacharparenright}\ auto\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} - in the primrec declaration; by default, the local name of - a class operation \isa{f} to instantiate on type constructor - \isa{{\isasymkappa}} are mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty, - these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command - or the corresponding ProofGeneral button.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Lifting and parametric types% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Overloaded definitions giving on class instantiation - may include recursion over the syntactic structure of types. - As a canonical example, we model product semigroups - using our simple algebra:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ fst\ p\isactrlisub {\isadigit{2}}{\isacharcomma}\ snd\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ snd\ p\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\isactrlisub {\isadigit{1}}\ p\isactrlisub {\isadigit{2}}\ p\isactrlisub {\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}\ {\isacharequal}\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\ \ \ \ \ \ \isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Associativity from product semigroups is - established using - the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical - associativity of the type components; these hypotheses - are facts due to the \isa{semigroup} constraints imposed - on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition. - Indeed, this pattern often occurs with parametric types - and type classes.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Subclassing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral) - by extending \isa{semigroup} - with one additional parameter \isa{neutral} together - with its property:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Again, we prove some instances, by - providing suitable parameter definitions and proofs for the - additional specifications. Observe that instantiations - for types with the same arity may be simultaneous:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ neutral{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ neutral{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isacharcomma}\ {\isasymone}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoidl\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoidl{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutl{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Fully-fledged monoids are modelled by another subclass - which does not add new parameters but tightens the specification:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline -\ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoid\ \isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\ \isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoid\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoid{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}p\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutr{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent To finish our small algebra example, we add a \isa{group} class - with a corresponding instance:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ inverse{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isacharequal}\ {\isacharminus}\ {\isacharparenleft}i{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ mult{\isacharunderscore}int{\isacharunderscore}def\ neutral{\isacharunderscore}int{\isacharunderscore}def\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsection{Type classes as locales% -} -\isamarkuptrue% -% -\isamarkupsubsection{A look behind the scene% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The example above gives an impression how Isar type classes work - in practice. As stated in the introduction, classes also provide - a link to Isar's locale system. Indeed, the logical core of a class - is nothing else than a locale:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ idem\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent essentially introduces the locale% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -\isanewline -% -\isadelimquote -\isanewline -% -\endisadelimquote -% -\isatagquote -\isacommand{locale}\isamarkupfalse% -\ idem\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent together with corresponding constant(s):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{consts}\isamarkupfalse% -\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The connection to the type system is done by means - of a primitive axclass% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -\isanewline -% -\isadelimquote -\isanewline -% -\endisadelimquote -% -\isatagquote -\isacommand{axclass}\isamarkupfalse% -\ idem\ {\isacharless}\ type\isanewline -\ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\begin{isamarkuptext}% -\noindent together with a corresponding interpretation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{interpretation}\isamarkupfalse% -\ idem{\isacharunderscore}class{\isacharcolon}\isanewline -\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\ \isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}rule\ idem{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This gives you at hand the full power of the Isabelle module system; - conclusions in locale \isa{idem} are implicitly propagated - to class \isa{idem}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsubsection{Abstract reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle locales enable reasoning at a general level, while results - are implicitly transferred to all instances. For example, we can - now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which - states that the function \isa{{\isacharparenleft}x\ {\isasymotimes}{\isacharparenright}} is injective:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% -\ assoc\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% -\ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target specification - indicates that the result is recorded within that context for later - use. This local theorem is also lifted to the global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of - \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Derived definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle locales support a concept of local definitions - in locales:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent If the locale \isa{group} is also a class, this local - definition is propagated onto a global definition of - \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} - with corresponding theorems - - \isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline% -pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}. - - \noindent As you can see from this example, for local - definitions you may use any specification tool - which works together with locales (e.g. \cite{krauss2006}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{A functor analogy% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We introduced Isar classes by analogy to type classes - functional programming; if we reconsider this in the - context of what has been said about type classes and locales, - we can drive this analogy further by stating that type - classes essentially correspond to functors which have - a canonical interpretation as type classes. - Anyway, there is also the possibility of other interpretations. - For example, also \isa{list}s form a monoid with - \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it - seems inappropriate to apply to lists - the same operations as for genuinely algebraic types. - In such a case, we simply can do a particular interpretation - of monoids for lists:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{interpretation}\isamarkupfalse% -\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \isacommand{qed}\isamarkupfalse% -\ auto% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This enables us to apply facts on monoids - to lists, e.g. \isa{{\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ x\ {\isacharequal}\ x}. - - When using this interpretation pattern, it may also - be appropriate to map derived definitions accordingly:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ replicate\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ list\ {\isasymRightarrow}\ {\isasymalpha}\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{interpretation}\isamarkupfalse% -\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{interpret}\isamarkupfalse% -\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ n\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ n\ {\isacharequal}\ replicate\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ n{\isacharparenright}\ auto\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\ intro{\isacharunderscore}locales% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsubsection{Additional subclass relations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Any \isa{group} is also a \isa{monoid}; this - can be made explicit by claiming an additional - subclass relation, - together with a proof of the logical difference:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{subclass}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ monoid\isanewline -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ invl\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}\ neutl\ invl\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ {\isasymone}{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ left{\isacharunderscore}cancel\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The logical proof is carried out on the locale level. - Afterwards it is propagated - to the type system, making \isa{group} an instance of - \isa{monoid} by adding an additional edge - to the graph of subclass relations - (cf.\ \figref{fig:subclass}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){\isa{semigroup}}} - \put(20,40){\makebox(0,0){\isa{monoidl}}} - \put(00,20){\makebox(0,0){\isa{monoid}}} - \put(40,00){\makebox(0,0){\isa{group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(25,35){\vector(1,-3){10}} - \end{picture} - \hspace{8em} - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){\isa{semigroup}}} - \put(20,40){\makebox(0,0){\isa{monoidl}}} - \put(00,20){\makebox(0,0){\isa{monoid}}} - \put(40,00){\makebox(0,0){\isa{group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(05,15){\vector(3,-1){30}} - \end{picture} - \caption{Subclass relationship of monoids and groups: - before and after establishing the relationship - \isa{group\ {\isasymsubseteq}\ monoid}; transitive edges left out.} - \label{fig:subclass} - \end{center} - \end{figure} -7 - For illustration, a derived definition - in \isa{group} which uses \isa{pow{\isacharunderscore}nat}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline -\ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline -\ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent yields the global definition of - \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} - with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{A note on syntax% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As a commodity, class context syntax allows to refer - to local class operations and their global counterparts - uniformly; type inference resolves ambiguities. For example:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{context}\isamarkupfalse% -\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{term}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % -\isamarkupcmt{example 1% -} -\isanewline -\isacommand{term}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % -\isamarkupcmt{example 2% -} -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{term}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % -\isamarkupcmt{example 3% -} -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Here in example 1, the term refers to the local class operation - \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint - enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}. - In the global context in example 3, the reference is - to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Further issues% -} -\isamarkuptrue% -% -\isamarkupsubsection{Type classes and code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Turning back to the first motivation for type classes, - namely overloading, it is obvious that overloading - stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and - \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} - targets naturally maps to Haskell type classes. - The code generator framework \cite{isabelle-codegen} - takes this into account. Concerning target languages - lacking type classes (e.g.~SML), type classes - are implemented by explicit dictionary construction. - As example, let's go back to the power function:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ example\ {\isacharcolon}{\isacharcolon}\ int\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This maps to Haskell as:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}module Example where {\char123}\\ -\hspace*{0pt}\\ -\hspace*{0pt}\\ -\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}nat{\char95}aux ::~Integer -> Nat -> Nat;\\ -\hspace*{0pt}nat{\char95}aux i n = (if i <= 0 then n else nat{\char95}aux (i - 1) (Suc n));\\ -\hspace*{0pt}\\ -\hspace*{0pt}nat ::~Integer -> Nat;\\ -\hspace*{0pt}nat i = nat{\char95}aux i Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}class Semigroup a where {\char123}\\ -\hspace*{0pt} ~mult ::~a -> a -> a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Semigroup a) => Monoidl a where {\char123}\\ -\hspace*{0pt} ~neutral ::~a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Monoidl a) => Monoid a where {\char123}\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Monoid a) => Group a where {\char123}\\ -\hspace*{0pt} ~inverse ::~a -> a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}inverse{\char95}int ::~Integer -> Integer;\\ -\hspace*{0pt}inverse{\char95}int i = negate i;\\ -\hspace*{0pt}\\ -\hspace*{0pt}neutral{\char95}int ::~Integer;\\ -\hspace*{0pt}neutral{\char95}int = 0;\\ -\hspace*{0pt}\\ -\hspace*{0pt}mult{\char95}int ::~Integer -> Integer -> Integer;\\ -\hspace*{0pt}mult{\char95}int i j = i + j;\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Semigroup Integer where {\char123}\\ -\hspace*{0pt} ~mult = mult{\char95}int;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Monoidl Integer where {\char123}\\ -\hspace*{0pt} ~neutral = neutral{\char95}int;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Monoid Integer where {\char123}\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Group Integer where {\char123}\\ -\hspace*{0pt} ~inverse = inverse{\char95}int;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\ -\hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\ -\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\ -\hspace*{0pt}\\ -\hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\ -\hspace*{0pt}pow{\char95}int k x =\\ -\hspace*{0pt} ~(if 0 <= k then pow{\char95}nat (nat k) x\\ -\hspace*{0pt} ~~~else inverse (pow{\char95}nat (nat (negate k)) x));\\ -\hspace*{0pt}\\ -\hspace*{0pt}example ::~Integer;\\ -\hspace*{0pt}example = pow{\char95}int 10 (-2);\\ -\hspace*{0pt}\\ -\hspace*{0pt}{\char125}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The whole code in SML with explicit dictionary passing:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun nat{\char95}aux i n =\\ -\hspace*{0pt} ~(if IntInf.<= (i,~(0 :~IntInf.int)) then n\\ -\hspace*{0pt} ~~~else nat{\char95}aux (IntInf.- (i,~(1 :~IntInf.int))) (Suc n));\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun nat i = nat{\char95}aux i Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ -\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a monoidl =\\ -\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl :~'a semigroup,~neutral :~'a{\char125};\\ -\hspace*{0pt}fun semigroup{\char95}monoidl (A{\char95}:'a monoidl) = {\char35}Classes{\char95}{\char95}semigroup{\char95}monoidl A{\char95};\\ -\hspace*{0pt}fun neutral (A{\char95}:'a monoidl) = {\char35}neutral A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a monoid = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid :~'a monoidl{\char125};\\ -\hspace*{0pt}fun monoidl{\char95}monoid (A{\char95}:'a monoid) = {\char35}Classes{\char95}{\char95}monoidl{\char95}monoid A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a group = {\char123}Classes{\char95}{\char95}monoid{\char95}group :~'a monoid,~inverse :~'a -> 'a{\char125};\\ -\hspace*{0pt}fun monoid{\char95}group (A{\char95}:'a group) = {\char35}Classes{\char95}{\char95}monoid{\char95}group A{\char95};\\ -\hspace*{0pt}fun inverse (A{\char95}:'a group) = {\char35}inverse A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ -\hspace*{0pt}\\ -\hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoidl{\char95}int =\\ -\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int monoidl;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoid{\char95}int = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int monoid;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val group{\char95}int =\\ -\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int group;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\ -\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\ -\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun pow{\char95}int A{\char95}~k x =\\ -\hspace*{0pt} ~(if IntInf.<= ((0 :~IntInf.int),~k)\\ -\hspace*{0pt} ~~~then pow{\char95}nat (monoid{\char95}group A{\char95}) (nat k) x\\ -\hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\ -\hspace*{0pt}\\ -\hspace*{0pt}val example :~IntInf.int =\\ -\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsubsection{Inspecting the type class universe% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -To facilitate orientation in complex subclass structures, - two diagnostics commands are provided: - - \begin{description} - - \item[\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isacharunderscore}classes}}}}] print a list of all classes - together with associated operations etc. - - \item[\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isacharunderscore}deps}}}}] visualizes the subclass relation - between all classes as a Hasse diagram. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/classes.tex --- a/doc-src/IsarAdvanced/Classes/classes.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,50 +0,0 @@ - -\documentclass[12pt,a4paper,fleqn]{report} -\usepackage{latexsym,graphicx} -\usepackage[refpage]{nomencl} -\usepackage{../../iman,../../extra,../../isar,../../proof} -\usepackage{../../isabelle,../../isabellesym} -\usepackage{style} -\usepackage{../../pdfsetup} - - -\hyphenation{Isabelle} -\hyphenation{Isar} -\isadroptag{theory} - -\title{\includegraphics[scale=0.5]{isabelle_isar} - \\[4ex] Haskell-style type classes with Isabelle/Isar} -\author{\emph{Florian Haftmann}} - -\begin{document} - -\maketitle - -\begin{abstract} - This tutorial introduces the look-and-feel of Isar type classes - to the end-user; Isar type classes are a convenient mechanism - for organizing specifications, overcoming some drawbacks - of raw axiomatic type classes. Essentially, they combine - an operational aspect (in the manner of Haskell) with - a logical aspect, both managed uniformly. -\end{abstract} - -\thispagestyle{empty}\clearpage - -\pagenumbering{roman} -\clearfirst - -\input{Thy/document/Classes.tex} - -\begingroup -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{../../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Classes/style.sty --- a/doc-src/IsarAdvanced/Classes/style.sty Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,43 +0,0 @@ - -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} -\newcommand{\figref}[1]{figure~\ref{#1}} - -%% logical markup -\newcommand{\strong}[1]{{\bfseries {#1}}} -\newcommand{\qn}[1]{\emph{#1}} - -%% typographic conventions -\newcommand{\qt}[1]{``{#1}''} - -%% verbatim text -\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} - -%% quoted segments -\makeatletter -\isakeeptag{quote} -\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}} -\renewcommand{\isatagquote}{\begin{quotesegment}} -\renewcommand{\endisatagquote}{\end{quotesegment}} -\makeatother - -%% presentation -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\binperiod -\underscoreoff - -\renewcommand{\isadigit}[1]{\isamath{#1}} - -\isabellestyle{it} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/IsaMakefile --- a/doc-src/IsarAdvanced/Codegen/IsaMakefile Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ - -## targets - -default: Thy -images: -test: Thy - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log - -USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document - - -## Thy - -THY = $(LOG)/HOL-Thy.gz - -Thy: $(THY) - -$(THY): Thy/ROOT.ML Thy/*.thy ../../antiquote_setup.ML - @$(USEDIR) HOL Thy - - -## clean - -clean: - @rm -f $(THY) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Makefile --- a/doc-src/IsarAdvanced/Codegen/Makefile Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = codegen - -FILES = $(NAME).tex Thy/document/*.tex \ - style.sty ../../iman.sty ../../extra.sty ../../isar.sty \ - ../../isabelle.sty ../../isabellesym.sty ../../pdfsetup.sty \ - ../../manual.bib ../../proof.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps codegen_process.ps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf codegen_process.pdf - $(PDFLATEX) $(NAME) - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,361 +0,0 @@ -theory Adaption -imports Setup -begin - -setup %invisible {* Code_Target.extend_target ("\", ("SML", K I)) *} - -section {* Adaption to target languages \label{sec:adaption} *} - -subsection {* Adapting code generation *} - -text {* - The aspects of code generation introduced so far have two aspects - in common: - - \begin{itemize} - \item They act uniformly, without reference to a specific - target language. - \item They are \emph{safe} in the sense that as long as you trust - the code generator meta theory and implementation, you cannot - produce programs that yield results which are not derivable - in the logic. - \end{itemize} - - \noindent In this section we will introduce means to \emph{adapt} the serialiser - to a specific target language, i.e.~to print program fragments - in a way which accommodates \qt{already existing} ingredients of - a target language environment, for three reasons: - - \begin{itemize} - \item improving readability and aesthetics of generated code - \item gaining efficiency - \item interface with language parts which have no direct counterpart - in @{text "HOL"} (say, imperative data structures) - \end{itemize} - - \noindent Generally, you should avoid using those features yourself - \emph{at any cost}: - - \begin{itemize} - \item The safe configuration methods act uniformly on every target language, - whereas for adaption you have to treat each target language separate. - \item Application is extremely tedious since there is no abstraction - which would allow for a static check, making it easy to produce garbage. - \item More or less subtle errors can be introduced unconsciously. - \end{itemize} - - \noindent However, even if you ought refrain from setting up adaption - yourself, already the @{text "HOL"} comes with some reasonable default - adaptions (say, using target language list syntax). There also some - common adaption cases which you can setup by importing particular - library theories. In order to understand these, we provide some clues here; - these however are not supposed to replace a careful study of the sources. -*} - -subsection {* The adaption principle *} - -text {* - The following figure illustrates what \qt{adaption} is conceptually - supposed to be: - - \begin{figure}[here] - \begin{tikzpicture}[scale = 0.5] - \tikzstyle water=[color = blue, thick] - \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] - \tikzstyle process=[color = green, semithick, ->] - \tikzstyle adaption=[color = red, semithick, ->] - \tikzstyle target=[color = black] - \foreach \x in {0, ..., 24} - \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin - + (0.25, -0.25) cos + (0.25, 0.25); - \draw[style=ice] (1, 0) -- - (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; - \draw[style=ice] (9, 0) -- - (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; - \draw[style=ice] (15, -6) -- - (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; - \draw[style=process] - (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); - \draw[style=process] - (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); - \node (adaption) at (11, -2) [style=adaption] {adaption}; - \node at (19, 3) [rotate=90] {generated}; - \node at (19.5, -5) {language}; - \node at (19.5, -3) {library}; - \node (includes) at (19.5, -1) {includes}; - \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 - \draw[style=process] - (includes) -- (serialisation); - \draw[style=process] - (reserved) -- (serialisation); - \draw[style=adaption] - (adaption) -- (serialisation); - \draw[style=adaption] - (adaption) -- (includes); - \draw[style=adaption] - (adaption) -- (reserved); - \end{tikzpicture} - \caption{The adaption principle} - \label{fig:adaption} - \end{figure} - - \noindent In the tame view, code generation acts as broker between - @{text logic}, @{text "intermediate language"} and - @{text "target language"} by means of @{text translation} and - @{text serialisation}; for the latter, the serialiser has to observe - the structure of the @{text language} itself plus some @{text reserved} - keywords which have to be avoided for generated code. - However, if you consider @{text adaption} mechanisms, the code generated - by the serializer is just the tip of the iceberg: - - \begin{itemize} - \item @{text serialisation} can be \emph{parametrised} such that - logical entities are mapped to target-specific ones - (e.g. target-specific list syntax, - see also \secref{sec:adaption_mechanisms}) - \item Such parametrisations can involve references to a - target-specific standard @{text library} (e.g. using - the @{text Haskell} @{verbatim Maybe} type instead - of the @{text HOL} @{type "option"} type); - if such are used, the corresponding identifiers - (in our example, @{verbatim Maybe}, @{verbatim Nothing} - and @{verbatim Just}) also have to be considered @{text reserved}. - \item Even more, the user can enrich the library of the - target-language by providing code snippets - (\qt{@{text "includes"}}) which are prepended to - any generated code (see \secref{sec:include}); this typically - also involves further @{text reserved} identifiers. - \end{itemize} - - \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms - have to act consistently; it is at the discretion of the user - to take care for this. -*} - -subsection {* Common adaption patterns *} - -text {* - The @{theory HOL} @{theory Main} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the @{text HOL} - library; beside being useful in applications, they may serve - as a tutorial for customising the code generator setup (see below - \secref{sec:adaption_mechanisms}). - - \begin{description} - - \item[@{theory "Code_Integer"}] represents @{text HOL} integers by big - integer literals in target languages. - \item[@{theory "Code_Char"}] represents @{text HOL} characters by - character literals in target languages. - \item[@{theory "Code_Char_chr"}] like @{text "Code_Char"}, - but also offers treatment of character codes; includes - @{theory "Code_Char"}. - \item[@{theory "Efficient_Nat"}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficiency; pattern - matching with @{term "0\nat"} / @{const "Suc"} - is eliminated; includes @{theory "Code_Integer"} - and @{theory "Code_Index"}. - \item[@{theory "Code_Index"}] provides an additional datatype - @{typ index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[@{theory "Code_Message"}] provides an additional datatype - @{typ message_string} which is isomorphic to strings; - @{typ message_string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn} -*} - - -subsection {* Parametrising serialisation \label{sec:adaption_mechanisms} *} - -text {* - Consider the following function and its corresponding - SML code: -*} - -primrec %quote in_interval :: "nat \ nat \ nat \ bool" where - "in_interval (k, l) n \ k \ n \ n \ l" -(*<*) -code_type %invisible bool - (SML) -code_const %invisible True and False and "op \" and Not - (SML and and and) -(*>*) -text %quote {*@{code_stmts in_interval (SML)}*} - -text {* - \noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialised as distinguished - entities with have nothing to do with the SML-built-in notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML @{verbatim "bool"} would be used. To map - the HOL @{typ bool} on SML @{verbatim "bool"}, we may use - \qn{custom serialisations}: -*} - -code_type %quotett bool - (SML "bool") -code_const %quotett True and False and "op \" - (SML "true" and "false" and "_ andalso _") - -text {* - \noindent The @{command code_type} command takes a type constructor - as arguments together with a list of custom serialisations. - Each custom serialisation starts with a target language - identifier followed by an expression, which during - code serialisation is inserted whenever the type constructor - would occur. For constants, @{command code_const} implements - the corresponding mechanism. Each ``@{verbatim "_"}'' in - a serialisation expression is treated as a placeholder - for the type constructor's (the constant's) arguments. -*} - -text %quote {*@{code_stmts in_interval (SML)}*} - -text {* - \noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serialiser - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here: -*} - -code_const %quotett "op \" - (SML infixl 1 "andalso") - -text %quote {*@{code_stmts in_interval (SML)}*} - -text {* - \noindent The attentive reader may ask how we assert that no generated - code will accidentally overwrite. For this reason the serialiser has - an internal table of identifiers which have to be avoided to be used - for new declarations. Initially, this table typically contains the - keywords of the target language. It can be extended manually, thus avoiding - accidental overwrites, using the @{command "code_reserved"} command: -*} - -code_reserved %quote "\" bool true false andalso - -text {* - \noindent Next, we try to map HOL pairs to SML pairs, using the - infix ``@{verbatim "*"}'' type constructor and parentheses: -*} -(*<*) -code_type %invisible * - (SML) -code_const %invisible Pair - (SML) -(*>*) -code_type %quotett * - (SML infix 2 "*") -code_const %quotett Pair - (SML "!((_),/ (_))") - -text {* - \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser - never to put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``@{verbatim "/"}'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serialisations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serialisations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialisation may be quoted - using ``@{verbatim "'"}''; thus, in - ``@{verbatim "fn '_ => _"}'' the first - ``@{verbatim "_"}'' is a proper underscore while the - second ``@{verbatim "_"}'' is a placeholder. -*} - - -subsection {* @{text Haskell} serialisation *} - -text {* - For convenience, the default - @{text HOL} setup for @{text Haskell} maps the @{class eq} class to - its counterpart in @{text Haskell}, giving custom serialisations - for the class @{class eq} (by command @{command code_class}) and its operation - @{const HOL.eq} -*} - -code_class %quotett eq - (Haskell "Eq") - -code_const %quotett "op =" - (Haskell infixl 4 "==") - -text {* - \noindent A problem now occurs whenever a type which - is an instance of @{class eq} in @{text HOL} is mapped - on a @{text Haskell}-built-in type which is also an instance - of @{text Haskell} @{text Eq}: -*} - -typedecl %quote bar - -instantiation %quote bar :: eq -begin - -definition %quote "eq_class.eq (x\bar) y \ x = y" - -instance %quote by default (simp add: eq_bar_def) - -end %quote - -code_type %quotett bar - (Haskell "Integer") - -text {* - \noindent The code generator would produce - an additional instance, which of course is rejected by the @{text Haskell} - compiler. - To suppress this additional instance, use - @{text "code_instance"}: -*} - -code_instance %quotett bar :: eq - (Haskell -) - - -subsection {* Enhancing the target language context \label{sec:include} *} - -text {* - In rare cases it is necessary to \emph{enrich} the context of a - target language; this is accomplished using the @{command "code_include"} - command: -*} - -code_include %quotett Haskell "Errno" -{*errno i = error ("Error number: " ++ show i)*} - -code_reserved %quotett Haskell Errno - -text {* - \noindent Such named @{text include}s are then prepended to every generated code. - Inspect such code in order to find out how @{command "code_include"} behaves - with respect to a particular target language. -*} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/Further.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Further.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,113 +0,0 @@ -theory Further -imports Setup -begin - -section {* Further issues \label{sec:further} *} - -subsection {* Further reading *} - -text {* - Do dive deeper into the issue of code generation, you should visit - the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which - contains exhaustive syntax diagrams. -*} - -subsection {* Modules *} - -text {* - When invoking the @{command export_code} command it is possible to leave - out the @{keyword "module_name"} part; then code is distributed over - different modules, where the module name space roughly is induced - by the @{text Isabelle} theory name space. - - Then sometimes the awkward situation occurs that dependencies between - definitions introduce cyclic dependencies between modules, which in the - @{text Haskell} world leaves you to the mercy of the @{text Haskell} implementation - you are using, while for @{text SML}/@{text OCaml} code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating -*} - -code_modulename %quote SML - A ABC - B ABC - C ABC - -text {* - we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialisation time. -*} - -subsection {* Evaluation oracle *} - -text {* - Code generation may also be used to \emph{evaluate} expressions - (using @{text SML} as target language of course). - For instance, the @{command value} allows to reduce an expression to a - normal form with respect to the underlying code equations: -*} - -value %quote "42 / (12 :: rat)" - -text {* - \noindent will display @{term "7 / (2 :: rat)"}. - - The @{method eval} method tries to reduce a goal by code generation to @{term True} - and solves it in that case, but fails otherwise: -*} - -lemma %quote "42 / (12 :: rat) = 7 / 2" - by %quote eval - -text {* - \noindent The soundness of the @{method eval} method depends crucially - on the correctness of the code generator; this is one of the reasons - why you should not use adaption (see \secref{sec:adaption}) frivolously. -*} - -subsection {* Code antiquotation *} - -text {* - In scenarios involving techniques like reflection it is quite common - that code generated from a theory forms the basis for implementing - a proof procedure in @{text SML}. To facilitate interfacing of generated code - with system code, the code generator provides a @{text code} antiquotation: -*} - -datatype %quote form = T | F | And form form | Or form form - -ML %quote {* - fun eval_form @{code T} = true - | eval_form @{code F} = false - | eval_form (@{code And} (p, q)) = - eval_form p andalso eval_form q - | eval_form (@{code Or} (p, q)) = - eval_form p orelse eval_form q; -*} - -text {* - \noindent @{text code} takes as argument the name of a constant; after the - whole @{text SML} is read, the necessary code is generated transparently - and the corresponding constant names are inserted. This technique also - allows to use pattern matching on constructors stemming from compiled - @{text datatypes}. - - For a less simplistic example, theory @{theory Ferrack} is - a good reference. -*} - -subsection {* Imperative data structures *} - -text {* - If you consider imperative data structures as inevitable for a specific - application, you should consider - \emph{Imperative Functional Programming with Isabelle/HOL} - (\cite{bulwahn-et-al:2008:imperative}); - the framework described there is available in theory @{theory Imperative_HOL}. -*} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,206 +0,0 @@ -theory Introduction -imports Setup -begin - -chapter {* Code generation from @{text "Isabelle/HOL"} theories *} - -section {* Introduction and Overview *} - -text {* - This tutorial introduces a generic code generator for the - @{text Isabelle} system. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and @{text Haskell} - \cite{haskell-revised-report}). - - Conceptually the code generator framework is part - of Isabelle's @{theory Pure} meta logic framework; the logic - @{theory HOL} which is an extension of @{theory Pure} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the @{theory HOL} distribution theories. - (see also \cite{isa-tutorial}). - - The code generator aims to be usable with no further ado - in most cases while allowing for detailed customisation. - This manifests in the structure of this tutorial: after a short - conceptual introduction with an example (\secref{sec:intro}), - we discuss the generic customisation facilities (\secref{sec:program}). - A further section (\secref{sec:adaption}) is dedicated to the matter of - \qn{adaption} to specific target language environments. After some - further issues (\secref{sec:further}) we conclude with an overview - of some ML programming interfaces (\secref{sec:ml}). - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the existing code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. In case of ambiguity, we will refer to the framework - described here as @{text "generic code generator"}, to the - other as @{text "SML code generator"}. - Also note that while the framework itself is - object-logic independent, only @{theory HOL} provides a reasonable - framework setup. - \end{warn} - -*} - -subsection {* Code generation via shallow embedding \label{sec:intro} *} - -text {* - The key concept for understanding @{text Isabelle}'s code generation is - \emph{shallow embedding}, i.e.~logical entities like constants, types and - classes are identified with corresponding concepts in the target language. - - Inside @{theory HOL}, the @{command datatype} and - @{command definition}/@{command primrec}/@{command fun} declarations form - the core of a functional programming language. The default code generator setup - allows to turn those into functional programs immediately. - This means that \qt{naive} code generation can proceed without further ado. - For example, here a simple \qt{implementation} of amortised queues: -*} - -datatype %quote 'a queue = AQueue "'a list" "'a list" - -definition %quote empty :: "'a queue" where - "empty = AQueue [] []" - -primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where - "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" - -fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where - "dequeue (AQueue [] []) = (None, AQueue [] [])" - | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" - | "dequeue (AQueue xs []) = - (case rev xs of y # ys \ (Some y, AQueue [] ys))" - -text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *} - -export_code %quote empty dequeue enqueue in SML - module_name Example file "examples/example.ML" - -text {* \noindent resulting in the following code: *} - -text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} - -text {* - \noindent The @{command export_code} command takes a space-separated list of - constants for which code shall be generated; anything else needed for those - is added implicitly. Then follows a target language identifier - (@{text SML}, @{text OCaml} or @{text Haskell}) and a freely chosen module name. - A file name denotes the destination to store the generated code. Note that - the semantics of the destination depends on the target language: for - @{text SML} and @{text OCaml} it denotes a \emph{file}, for @{text Haskell} - it denotes a \emph{directory} where a file named as the module name - (with extension @{text ".hs"}) is written: -*} - -export_code %quote empty dequeue enqueue in Haskell - module_name Example file "examples/" - -text {* - \noindent This is how the corresponding code in @{text Haskell} looks like: -*} - -text %quote {*@{code_stmts empty enqueue dequeue (Haskell)}*} - -text {* - \noindent This demonstrates the basic usage of the @{command export_code} command; - for more details see \secref{sec:further}. -*} - -subsection {* Code generator architecture \label{sec:concept} *} - -text {* - What you have seen so far should be already enough in a lot of cases. If you - are content with this, you can quit reading here. Anyway, in order to customise - and adapt the code generator, it is inevitable to gain some understanding - how it works. - - \begin{figure}[h] - \begin{tikzpicture}[x = 4.2cm, y = 1cm] - \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; - \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; - \tikzstyle process_arrow=[->, semithick, color = green]; - \node (HOL) at (0, 4) [style=entity] {@{text "Isabelle/HOL"} theory}; - \node (eqn) at (2, 2) [style=entity] {code equations}; - \node (iml) at (2, 0) [style=entity] {intermediate language}; - \node (seri) at (1, 0) [style=process] {serialisation}; - \node (SML) at (0, 3) [style=entity] {@{text SML}}; - \node (OCaml) at (0, 2) [style=entity] {@{text OCaml}}; - \node (further) at (0, 1) [style=entity] {@{text "\"}}; - \node (Haskell) at (0, 0) [style=entity] {@{text Haskell}}; - \draw [style=process_arrow] (HOL) .. controls (2, 4) .. - node [style=process, near start] {selection} - node [style=process, near end] {preprocessing} - (eqn); - \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); - \draw [style=process_arrow] (iml) -- (seri); - \draw [style=process_arrow] (seri) -- (SML); - \draw [style=process_arrow] (seri) -- (OCaml); - \draw [style=process_arrow, dashed] (seri) -- (further); - \draw [style=process_arrow] (seri) -- (Haskell); - \end{tikzpicture} - \caption{Code generator architecture} - \label{fig:arch} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{code equations}, \emph{datatypes}, and - \emph{type classes}. A code equation as a first approximation - is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \ t\<^isub>n \ t"} - (an equation headed by a constant @{text f} with arguments - @{text "t\<^isub>1 t\<^isub>2 \ t\<^isub>n"} and right hand side @{text t}). - Code generation aims to turn code equations - into a functional program. This is achieved by three major - components which operate sequentially, i.e. the result of one is - the input - of the next in the chain, see diagram \ref{fig:arch}: - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modelling - code equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item Before the selected code equations are continued with, - they can be \qn{preprocessed}, i.e. subjected to theorem - transformations. This \qn{preprocessor} is an interface which - allows to apply - the full expressiveness of ML-based theorem transformations - to code generation; motivating examples are shown below, see - \secref{sec:preproc}. - The result of the preprocessing step is a structured collection - of code equations. - - \item These code equations are \qn{translated} to a program - in an abstract intermediate language. Think of it as a kind - of \qt{Mini-Haskell} with four \qn{statements}: @{text data} - (for datatypes), @{text fun} (stemming from code equations), - also @{text class} and @{text inst} (for type classes). - - \item Finally, the abstract program is \qn{serialised} into concrete - source code of a target language. - - \end{itemize} - - \noindent From these steps, only the two last are carried out outside the logic; by - keeping this layer as thin as possible, the amount of code to trust is - kept to a minimum. -*} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/ML.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/ML.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,177 +0,0 @@ -theory "ML" -imports Setup -begin - -section {* ML system interfaces \label{sec:ml} *} - -text {* - Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces. -*} - -subsection {* Executable theory content: @{text Code} *} - -text {* - This Pure module implements the core notions of - executable content of a theory. -*} - -subsubsection {* Managing executable content *} - -text %mlref {* - \begin{mldecls} - @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ - @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ - @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\ - @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\ - @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\ - @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option) - -> theory -> theory"} \\ - @{index_ML Code.del_functrans: "string -> theory -> theory"} \\ - @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ - @{index_ML Code.get_datatype: "theory -> string - -> (string * sort) list * (string * typ list) list"} \\ - @{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"} - \end{mldecls} - - \begin{description} - - \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function - theorem @{text "thm"} to executable content. - - \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function - theorem @{text "thm"} from executable content, if present. - - \item @{ML Code.add_eqnl}~@{text "(const, lthms)"}~@{text "thy"} adds - suspended code equations @{text lthms} for constant - @{text const} to executable content. - - \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes - the preprocessor simpset. - - \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds - function transformer @{text f} (named @{text name}) to executable content; - @{text f} is a transformer of the code equations belonging - to a certain function definition, depending on the - current theory context. Returning @{text NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new code equations. - - \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes - function transformer named @{text name} from executable content. - - \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds - a datatype to executable content, with generation - set @{text cs}. - - \item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"} - returns type constructor corresponding to - constructor @{text const}; returns @{text NONE} - if @{text const} is no constructor. - - \end{description} -*} - -subsection {* Auxiliary *} - -text %mlref {* - \begin{mldecls} - @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\ - @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\ - @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML Code_Unit.read_const}~@{text thy}~@{text s} - reads a constant as a concrete term expression @{text s}. - - \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm} - extracts the constant and its type from a code equation @{text thm}. - - \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm} - rewrites a code equation @{text thm} with a simpset @{text ss}; - only arguments and right hand side are rewritten, - not the head of the code equation. - - \end{description} - -*} - -subsection {* Implementing code generator applications *} - -text {* - Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things. -*} - -subsubsection {* Data depending on the theory's executable content *} - -text {* - Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot @{ML_functor CodeDataFun}; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - @{text "type T"} \\ - @{text "val empty: T"} \\ - @{text "val purge: theory \ string list option \ T \ T"} - \end{tabular} - - \begin{description} - - \item @{text T} the type of data to store. - - \item @{text empty} initial (empty) data. - - \item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content; - @{text consts} indicates the kind - of change: @{ML NONE} stands for a fundamental change - which invalidates any existing code, @{text "SOME consts"} - hints that executable content for constants @{text consts} - has changed. - - \end{description} - - \noindent An instance of @{ML_functor CodeDataFun} provides the following - interface: - - \medskip - \begin{tabular}{l} - @{text "get: theory \ T"} \\ - @{text "change: theory \ (T \ T) \ T"} \\ - @{text "change_yield: theory \ (T \ 'a * T) \ 'a * T"} - \end{tabular} - - \begin{description} - - \item @{text get} retrieval of the current data. - - \item @{text change} update of current data (cached!) - by giving a continuation. - - \item @{text change_yield} update with side result. - - \end{description} -*} - -text {* - \bigskip - - \emph{Happy proving, happy hacking!} -*} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/Program.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Program.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,534 +0,0 @@ -theory Program -imports Introduction -begin - -section {* Turning Theories into Programs \label{sec:program} *} - -subsection {* The @{text "Isabelle/HOL"} default setup *} - -text {* - We have already seen how by default equations stemming from - @{command definition}/@{command primrec}/@{command fun} - statements are used for code generation. This default behaviour - can be changed, e.g. by providing different code equations. - All kinds of customisation shown in this section is \emph{safe} - in the sense that the user does not have to worry about - correctness -- all programs generatable that way are partially - correct. -*} - -subsection {* Selecting code equations *} - -text {* - Coming back to our introductory example, we - could provide an alternative code equations for @{const dequeue} - explicitly: -*} - -lemma %quote [code]: - "dequeue (AQueue xs []) = - (if xs = [] then (None, AQueue [] []) - else dequeue (AQueue [] (rev xs)))" - "dequeue (AQueue xs (y # ys)) = - (Some y, AQueue xs ys)" - by (cases xs, simp_all) (cases "rev xs", simp_all) - -text {* - \noindent The annotation @{text "[code]"} is an @{text Isar} - @{text attribute} which states that the given theorems should be - considered as code equations for a @{text fun} statement -- - the corresponding constant is determined syntactically. The resulting code: -*} - -text %quote {*@{code_stmts dequeue (consts) dequeue (Haskell)}*} - -text {* - \noindent You may note that the equality test @{term "xs = []"} has been - replaced by the predicate @{term "null xs"}. This is due to the default - setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). - - Changing the default constructor set of datatypes is also - possible. See \secref{sec:datatypes} for an example. - - As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the @{command code_thms} command: -*} - -code_thms %quote dequeue - -text {* - \noindent prints a table with \emph{all} code equations - for @{const dequeue}, including - \emph{all} code equations those equations depend - on recursively. - - Similarly, the @{command code_deps} command shows a graph - visualising dependencies between code equations. -*} - -subsection {* @{text class} and @{text instantiation} *} - -text {* - Concerning type classes and code generation, let us examine an example - from abstract algebra: -*} - -class %quote semigroup = - fixes mult :: "'a \ 'a \ 'a" (infixl "\" 70) - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - -class %quote monoid = semigroup + - fixes neutral :: 'a ("\") - assumes neutl: "\ \ x = x" - and neutr: "x \ \ = x" - -instantiation %quote nat :: monoid -begin - -primrec %quote mult_nat where - "0 \ n = (0\nat)" - | "Suc m \ n = n + m \ n" - -definition %quote neutral_nat where - "\ = Suc 0" - -lemma %quote add_mult_distrib: - fixes n m q :: nat - shows "(n + m) \ q = n \ q + m \ q" - by (induct n) simp_all - -instance %quote proof - fix m n q :: nat - show "m \ n \ q = m \ (n \ q)" - by (induct m) (simp_all add: add_mult_distrib) - show "\ \ n = n" - by (simp add: neutral_nat_def) - show "m \ \ = m" - by (induct m) (simp_all add: neutral_nat_def) -qed - -end %quote - -text {* - \noindent We define the natural operation of the natural numbers - on monoids: -*} - -primrec %quote (in monoid) pow :: "nat \ 'a \ 'a" where - "pow 0 a = \" - | "pow (Suc n) a = a \ pow n a" - -text {* - \noindent This we use to define the discrete exponentiation function: -*} - -definition %quote bexp :: "nat \ nat" where - "bexp n = pow n (Suc (Suc 0))" - -text {* - \noindent The corresponding code: -*} - -text %quote {*@{code_stmts bexp (Haskell)}*} - -text {* - \noindent This is a convenient place to show how explicit dictionary construction - manifests in generated code (here, the same example in @{text SML}): -*} - -text %quote {*@{code_stmts bexp (SML)}*} - -text {* - \noindent Note the parameters with trailing underscore (@{verbatim "A_"}) - which are the dictionary parameters. -*} - -subsection {* The preprocessor \label{sec:preproc} *} - -text {* - Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as code equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared and undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - - Some common applications: -*} - -text_raw {* - \begin{itemize} -*} - -text {* - \item replacing non-executable constructs by executable ones: -*} - -lemma %quote [code inline]: - "x \ set xs \ x mem xs" by (induct xs) simp_all - -text {* - \item eliminating superfluous constants: -*} - -lemma %quote [code inline]: - "1 = Suc 0" by simp - -text {* - \item replacing executable but inconvenient constructs: -*} - -lemma %quote [code inline]: - "xs = [] \ List.null xs" by (induct xs) simp_all - -text_raw {* - \end{itemize} -*} - -text {* - \noindent \emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The @{term "0\nat"} / @{const Suc} - pattern elimination implemented in - theory @{text Efficient_Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the @{command print_codesetup} command. - @{command code_thms} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on code equations. - - \begin{warn} - The attribute \emph{code unfold} - associated with the @{text "SML code generator"} also applies to - the @{text "generic code generator"}: - \emph{code unfold} implies \emph{code inline}. - \end{warn} -*} - -subsection {* Datatypes \label{sec:datatypes} *} - -text {* - Conceptually, any datatype is spanned by a set of - \emph{constructors} of type @{text "\ = \ \ \ \\<^isub>1 \ \\<^isub>n"} where @{text - "{\\<^isub>1, \, \\<^isub>n}"} is exactly the set of \emph{all} type variables in - @{text "\"}. The HOL datatype package by default registers any new - datatype in the table of datatypes, which may be inspected using the - @{command print_codesetup} command. - - In some cases, it is appropriate to alter or extend this table. As - an example, we will develop an alternative representation of the - queue example given in \secref{sec:intro}. The amortised - representation is convenient for generating code but exposes its - \qt{implementation} details, which may be cumbersome when proving - theorems about it. Therefore, here a simple, straightforward - representation of queues: -*} - -datatype %quote 'a queue = Queue "'a list" - -definition %quote empty :: "'a queue" where - "empty = Queue []" - -primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where - "enqueue x (Queue xs) = Queue (xs @ [x])" - -fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where - "dequeue (Queue []) = (None, Queue [])" - | "dequeue (Queue (x # xs)) = (Some x, Queue xs)" - -text {* - \noindent This we can use directly for proving; for executing, - we provide an alternative characterisation: -*} - -definition %quote AQueue :: "'a list \ 'a list \ 'a queue" where - "AQueue xs ys = Queue (ys @ rev xs)" - -code_datatype %quote AQueue - -text {* - \noindent Here we define a \qt{constructor} @{const "AQueue"} which - is defined in terms of @{text "Queue"} and interprets its arguments - according to what the \emph{content} of an amortised queue is supposed - to be. Equipped with this, we are able to prove the following equations - for our primitive queue operations which \qt{implement} the simple - queues in an amortised fashion: -*} - -lemma %quote empty_AQueue [code]: - "empty = AQueue [] []" - unfolding AQueue_def empty_def by simp - -lemma %quote enqueue_AQueue [code]: - "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" - unfolding AQueue_def by simp - -lemma %quote dequeue_AQueue [code]: - "dequeue (AQueue xs []) = - (if xs = [] then (None, AQueue [] []) - else dequeue (AQueue [] (rev xs)))" - "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" - unfolding AQueue_def by simp_all - -text {* - \noindent For completeness, we provide a substitute for the - @{text case} combinator on queues: -*} - -definition %quote - aqueue_case_def: "aqueue_case = queue_case" - -lemma %quote aqueue_case [code, code inline]: - "queue_case = aqueue_case" - unfolding aqueue_case_def .. - -lemma %quote case_AQueue [code]: - "aqueue_case f (AQueue xs ys) = f (ys @ rev xs)" - unfolding aqueue_case_def AQueue_def by simp - -text {* - \noindent The resulting code looks as expected: -*} - -text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} - -text {* - \noindent From this example, it can be glimpsed that using own - constructor sets is a little delicate since it changes the set of - valid patterns for values of that type. Without going into much - detail, here some practical hints: - - \begin{itemize} - - \item When changing the constructor set for datatypes, take care - to provide an alternative for the @{text case} combinator - (e.g.~by replacing it using the preprocessor). - - \item Values in the target language need not to be normalised -- - different values in the target language may represent the same - value in the logic. - - \item Usually, a good methodology to deal with the subtleties of - pattern matching is to see the type as an abstract type: provide - a set of operations which operate on the concrete representation - of the type, and derive further operations by combinations of - these primitive ones, without relying on a particular - representation. - - \end{itemize} -*} - - -subsection {* Equality and wellsortedness *} - -text {* - Surely you have already noticed how equality is treated - by the code generator: -*} - -primrec %quote collect_duplicates :: "'a list \ 'a list \ 'a list \ 'a list" where - "collect_duplicates xs ys [] = xs" - | "collect_duplicates xs ys (z#zs) = (if z \ set xs - then if z \ set ys - then collect_duplicates xs ys zs - else collect_duplicates xs (z#ys) zs - else collect_duplicates (z#xs) (z#ys) zs)" - -text {* - \noindent The membership test during preprocessing is rewritten, - resulting in @{const List.member}, which itself - performs an explicit equality check. -*} - -text %quote {*@{code_stmts collect_duplicates (SML)}*} - -text {* - \noindent Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class @{class eq} with a corresponding operation - @{const eq_class.eq} such that @{thm eq [no_vars]}. - The preprocessing framework does the rest by propagating the - @{class eq} constraints through all dependent code equations. - For datatypes, instances of @{class eq} are implicitly derived - when possible. For other types, you may instantiate @{text eq} - manually like any other type class. - - Though this @{text eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples - (also see theory @{theory Product_ord}): -*} - -instantiation %quote "*" :: (order, order) order -begin - -definition %quote [code del]: - "x \ y \ fst x < fst y \ fst x = fst y \ snd x \ snd y" - -definition %quote [code del]: - "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" - -instance %quote proof -qed (auto simp: less_eq_prod_def less_prod_def intro: order_less_trans) - -end %quote - -lemma %quote order_prod [code]: - "(x1 \ 'a\order, y1 \ 'b\order) < (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 < y2" - "(x1 \ 'a\order, y1 \ 'b\order) \ (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 \ y2" - by (simp_all add: less_prod_def less_eq_prod_def) - -text {* - \noindent Then code generation will fail. Why? The definition - of @{term "op \"} depends on equality on both arguments, - which are polymorphic and impose an additional @{class eq} - class constraint, which the preprocessor does not propagate - (for technical reasons). - - The solution is to add @{class eq} explicitly to the first sort arguments in the - code theorems: -*} - -lemma %quote order_prod_code [code]: - "(x1 \ 'a\{order, eq}, y1 \ 'b\order) < (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 < y2" - "(x1 \ 'a\{order, eq}, y1 \ 'b\order) \ (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 \ y2" - by (simp_all add: less_prod_def less_eq_prod_def) - -text {* - \noindent Then code generation succeeds: -*} - -text %quote {*@{code_stmts "op \ \ _ \ _ \ _ \ _ \ bool" (SML)}*} - -text {* - In some cases, the automatically derived code equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers): -*} - -datatype %quote monotype = Mono nat "monotype list" -(*<*) -lemma monotype_eq: - "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ - eq_class.eq tyco1 tyco2 \ eq_class.eq typargs1 typargs2" by (simp add: eq) -(*>*) - -text {* - \noindent Then code generation for SML would fail with a message - that the generated code contains illegal mutual dependencies: - the theorem @{thm monotype_eq [no_vars]} already requires the - instance @{text "monotype \ eq"}, which itself requires - @{thm monotype_eq [no_vars]}; Haskell has no problem with mutually - recursive @{text instance} and @{text function} definitions, - but the SML serialiser does not support this. - - In such cases, you have to provide your own equality equations - involving auxiliary constants. In our case, - @{const [show_types] list_all2} can do the job: -*} - -lemma %quote monotype_eq_list_all2 [code]: - "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ - eq_class.eq tyco1 tyco2 \ list_all2 eq_class.eq typargs1 typargs2" - by (simp add: eq list_all2_eq [symmetric]) - -text {* - \noindent does not depend on instance @{text "monotype \ eq"}: -*} - -text %quote {*@{code_stmts "eq_class.eq :: monotype \ monotype \ bool" (SML)}*} - - -subsection {* Explicit partiality *} - -text {* - Partiality usually enters the game by partial patterns, as - in the following example, again for amortised queues: -*} - -definition %quote strict_dequeue :: "'a queue \ 'a \ 'a queue" where - "strict_dequeue q = (case dequeue q - of (Some x, q') \ (x, q'))" - -lemma %quote strict_dequeue_AQueue [code]: - "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)" - "strict_dequeue (AQueue xs []) = - (case rev xs of y # ys \ (y, AQueue [] ys))" - by (simp_all add: strict_dequeue_def dequeue_AQueue split: list.splits) - -text {* - \noindent In the corresponding code, there is no equation - for the pattern @{term "AQueue [] []"}: -*} - -text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*} - -text {* - \noindent In some cases it is desirable to have this - pseudo-\qt{partiality} more explicitly, e.g.~as follows: -*} - -axiomatization %quote empty_queue :: 'a - -definition %quote strict_dequeue' :: "'a queue \ 'a \ 'a queue" where - "strict_dequeue' q = (case dequeue q of (Some x, q') \ (x, q') | _ \ empty_queue)" - -lemma %quote strict_dequeue'_AQueue [code]: - "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue - else strict_dequeue' (AQueue [] (rev xs)))" - "strict_dequeue' (AQueue xs (y # ys)) = - (y, AQueue xs ys)" - by (simp_all add: strict_dequeue'_def dequeue_AQueue split: list.splits) - -text {* - Observe that on the right hand side of the definition of @{const - "strict_dequeue'"} the constant @{const empty_queue} occurs - which is unspecified. - - Normally, if constants without any code equations occur in a - program, the code generator complains (since in most cases this is - not what the user expects). But such constants can also be thought - of as function definitions with no equations which always fail, - since there is never a successful pattern match on the left hand - side. In order to categorise a constant into that category - explicitly, use @{command "code_abort"}: -*} - -code_abort %quote empty_queue - -text {* - \noindent Then the code generator will just insert an error or - exception at the appropriate position: -*} - -text %quote {*@{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}*} - -text {* - \noindent This feature however is rarely needed in practice. - Note also that the @{text HOL} default setup already declares - @{const undefined} as @{command "code_abort"}, which is most - likely to be used in such situations. -*} - -end - \ No newline at end of file diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/ROOT.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/ROOT.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,11 +0,0 @@ - -(* $Id$ *) - -no_document use_thy "Setup"; -no_document use_thys ["Efficient_Nat"]; - -use_thy "Introduction"; -use_thy "Program"; -use_thy "Adaption"; -use_thy "Further"; -use_thy "ML"; diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/Setup.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,12 +0,0 @@ -theory Setup -imports Complex_Main -uses "../../../antiquote_setup.ML" "../../../more_antiquote.ML" -begin - -ML {* no_document use_thys - ["Efficient_Nat", "Code_Char_chr", "Product_ord", "~~/src/HOL/Imperative_HOL/Imperative_HOL", - "~~/src/HOL/Decision_Procs/Ferrack"] *} - -ML_val {* Code_Target.code_width := 74 *} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,679 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Adaption}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Adaption\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Code{\isacharunderscore}Target{\isachardot}extend{\isacharunderscore}target\ {\isacharparenleft}{\isachardoublequote}{\isasymSML}{\isachardoublequote}{\isacharcomma}\ {\isacharparenleft}{\isachardoublequote}SML{\isachardoublequote}{\isacharcomma}\ K\ I{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsection{Adaption to target languages \label{sec:adaption}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Adapting code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The aspects of code generation introduced so far have two aspects - in common: - - \begin{itemize} - \item They act uniformly, without reference to a specific - target language. - \item They are \emph{safe} in the sense that as long as you trust - the code generator meta theory and implementation, you cannot - produce programs that yield results which are not derivable - in the logic. - \end{itemize} - - \noindent In this section we will introduce means to \emph{adapt} the serialiser - to a specific target language, i.e.~to print program fragments - in a way which accommodates \qt{already existing} ingredients of - a target language environment, for three reasons: - - \begin{itemize} - \item improving readability and aesthetics of generated code - \item gaining efficiency - \item interface with language parts which have no direct counterpart - in \isa{HOL} (say, imperative data structures) - \end{itemize} - - \noindent Generally, you should avoid using those features yourself - \emph{at any cost}: - - \begin{itemize} - \item The safe configuration methods act uniformly on every target language, - whereas for adaption you have to treat each target language separate. - \item Application is extremely tedious since there is no abstraction - which would allow for a static check, making it easy to produce garbage. - \item More or less subtle errors can be introduced unconsciously. - \end{itemize} - - \noindent However, even if you ought refrain from setting up adaption - yourself, already the \isa{HOL} comes with some reasonable default - adaptions (say, using target language list syntax). There also some - common adaption cases which you can setup by importing particular - library theories. In order to understand these, we provide some clues here; - these however are not supposed to replace a careful study of the sources.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The adaption principle% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following figure illustrates what \qt{adaption} is conceptually - supposed to be: - - \begin{figure}[here] - \begin{tikzpicture}[scale = 0.5] - \tikzstyle water=[color = blue, thick] - \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] - \tikzstyle process=[color = green, semithick, ->] - \tikzstyle adaption=[color = red, semithick, ->] - \tikzstyle target=[color = black] - \foreach \x in {0, ..., 24} - \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin - + (0.25, -0.25) cos + (0.25, 0.25); - \draw[style=ice] (1, 0) -- - (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; - \draw[style=ice] (9, 0) -- - (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; - \draw[style=ice] (15, -6) -- - (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; - \draw[style=process] - (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); - \draw[style=process] - (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); - \node (adaption) at (11, -2) [style=adaption] {adaption}; - \node at (19, 3) [rotate=90] {generated}; - \node at (19.5, -5) {language}; - \node at (19.5, -3) {library}; - \node (includes) at (19.5, -1) {includes}; - \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 - \draw[style=process] - (includes) -- (serialisation); - \draw[style=process] - (reserved) -- (serialisation); - \draw[style=adaption] - (adaption) -- (serialisation); - \draw[style=adaption] - (adaption) -- (includes); - \draw[style=adaption] - (adaption) -- (reserved); - \end{tikzpicture} - \caption{The adaption principle} - \label{fig:adaption} - \end{figure} - - \noindent In the tame view, code generation acts as broker between - \isa{logic}, \isa{intermediate\ language} and - \isa{target\ language} by means of \isa{translation} and - \isa{serialisation}; for the latter, the serialiser has to observe - the structure of the \isa{language} itself plus some \isa{reserved} - keywords which have to be avoided for generated code. - However, if you consider \isa{adaption} mechanisms, the code generated - by the serializer is just the tip of the iceberg: - - \begin{itemize} - \item \isa{serialisation} can be \emph{parametrised} such that - logical entities are mapped to target-specific ones - (e.g. target-specific list syntax, - see also \secref{sec:adaption_mechanisms}) - \item Such parametrisations can involve references to a - target-specific standard \isa{library} (e.g. using - the \isa{Haskell} \verb|Maybe| type instead - of the \isa{HOL} \isa{option} type); - if such are used, the corresponding identifiers - (in our example, \verb|Maybe|, \verb|Nothing| - and \verb|Just|) also have to be considered \isa{reserved}. - \item Even more, the user can enrich the library of the - target-language by providing code snippets - (\qt{\isa{includes}}) which are prepended to - any generated code (see \secref{sec:include}); this typically - also involves further \isa{reserved} identifiers. - \end{itemize} - - \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms - have to act consistently; it is at the discretion of the user - to take care for this.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Common adaption patterns% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \hyperlink{theory.Main}{\mbox{\isa{Main}}} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the \isa{HOL} - library; beside being useful in applications, they may serve - as a tutorial for customising the code generator setup (see below - \secref{sec:adaption_mechanisms}). - - \begin{description} - - \item[\hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}}] represents \isa{HOL} integers by big - integer literals in target languages. - \item[\hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}] represents \isa{HOL} characters by - character literals in target languages. - \item[\hyperlink{theory.Code-Char-chr}{\mbox{\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}}}] like \isa{Code{\isacharunderscore}Char}, - but also offers treatment of character codes; includes - \hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}. - \item[\hyperlink{theory.Efficient-Nat}{\mbox{\isa{Efficient{\isacharunderscore}Nat}}}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficiency; pattern - matching with \isa{{\isadigit{0}}} / \isa{Suc} - is eliminated; includes \hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}} - and \hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}. - \item[\hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}] provides an additional datatype - \isa{index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[\hyperlink{theory.Code-Message}{\mbox{\isa{Code{\isacharunderscore}Message}}}] provides an additional datatype - \isa{message{\isacharunderscore}string} which is isomorphic to strings; - \isa{message{\isacharunderscore}string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Parametrising serialisation \label{sec:adaption_mechanisms}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function and its corresponding - SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype boola = True | False;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun anda x True = x\\ -\hspace*{0pt} ~| anda x False = False\\ -\hspace*{0pt} ~| anda True x = x\\ -\hspace*{0pt} ~| anda False x = False;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = False\\ -\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialised as distinguished - entities with have nothing to do with the SML-built-in notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML \verb|bool| would be used. To map - the HOL \isa{bool} on SML \verb|bool|, we may use - \qn{custom serialisations}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bool\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent The \hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} command takes a type constructor - as arguments together with a list of custom serialisations. - Each custom serialisation starts with a target language - identifier followed by an expression, which during - code serialisation is inserted whenever the type constructor - would occur. For constants, \hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} implements - the corresponding mechanism. Each ``\verb|_|'' in - a serialisation expression is treated as a placeholder - for the type constructor's (the constant's) arguments.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ -\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun in{\char95}interval (k,~l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serialiser - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ -\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun in{\char95}interval (k,~l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The attentive reader may ask how we assert that no generated - code will accidentally overwrite. For this reason the serialiser has - an internal table of identifiers which have to be avoided to be used - for new declarations. Initially, this table typically contains the - keywords of the target language. It can be extended manually, thus avoiding - accidental overwrites, using the \hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymSML}{\isachardoublequoteclose}\ bool\ true\ false\ andalso% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Next, we try to map HOL pairs to SML pairs, using the - infix ``\verb|*|'' type constructor and parentheses:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ {\isacharasterisk}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ Pair\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent The initial bang ``\verb|!|'' tells the serialiser - never to put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``\verb|/|'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serialisations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serialisations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialisation may be quoted - using ``\verb|'|''; thus, in - ``\verb|fn '_ => _|'' the first - ``\verb|_|'' is a proper underscore while the - second ``\verb|_|'' is a placeholder.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{\isa{Haskell} serialisation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For convenience, the default - \isa{HOL} setup for \isa{Haskell} maps the \isa{eq} class to - its counterpart in \isa{Haskell}, giving custom serialisations - for the class \isa{eq} (by command \hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}) and its operation - \isa{eq{\isacharunderscore}class{\isachardot}eq}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}class}\isamarkupfalse% -\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent A problem now occurs whenever a type which - is an instance of \isa{eq} in \isa{HOL} is mapped - on a \isa{Haskell}-built-in type which is also an instance - of \isa{Haskell} \isa{Eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{typedecl}\isamarkupfalse% -\ bar\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{by}\isamarkupfalse% -\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -\isanewline -% -\isadelimquotett -\isanewline -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bar\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent The code generator would produce - an additional instance, which of course is rejected by the \isa{Haskell} - compiler. - To suppress this additional instance, use - \isa{code{\isacharunderscore}instance}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\isamarkupsubsection{Enhancing the target language context \label{sec:include}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In rare cases it is necessary to \emph{enrich} the context of a - target language; this is accomplished using the \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} - command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}include}\isamarkupfalse% -\ Haskell\ {\isachardoublequoteopen}Errno{\isachardoublequoteclose}\isanewline -{\isacharverbatimopen}errno\ i\ {\isacharequal}\ error\ {\isacharparenleft}{\isachardoublequote}Error\ number{\isacharcolon}\ {\isachardoublequote}\ {\isacharplus}{\isacharplus}\ show\ i{\isacharparenright}{\isacharverbatimclose}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% -\ Haskell\ Errno% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent Such named \isa{include}s are then prepended to every generated code. - Inspect such code in order to find out how \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} behaves - with respect to a particular target language.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1690 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Codegen}% -% -\isadelimtheory -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isamarkupchapter{Code generation from Isabelle theories% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction% -} -\isamarkuptrue% -% -\isamarkupsubsection{Motivation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Executing formal specifications as programs is a well-established - topic in the theorem proving community. With increasing - application of theorem proving systems in the area of - software development and verification, its relevance manifests - for running test cases and rapid prototyping. In logical - calculi like constructive type theory, - a notion of executability is implicit due to the nature - of the calculus. In contrast, specifications in Isabelle - can be highly non-executable. In order to bridge - the gap between logic and executable specifications, - an explicit non-trivial transformation has to be applied: - code generation. - - This tutorial introduces a generic code generator for the - Isabelle system \cite{isa-tutorial}. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell - \cite{haskell-revised-report}). - We aim to provide a - versatile environment - suitable for software development and verification, - structuring the process - of code generation into a small set of orthogonal principles - while achieving a big coverage of application areas - with maximum flexibility. - - Conceptually the code generator framework is part - of Isabelle's \isa{Pure} meta logic; the object logic - \isa{HOL} which is an extension of \isa{Pure} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the \isa{HOL} \emph{Main} theory - (see also \cite{isa-tutorial}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The code generator aims to be usable with no further ado - in most cases while allowing for detailed customization. - This manifests in the structure of this tutorial: - we start with a generic example \secref{sec:example} - and introduce code generation concepts \secref{sec:concept}. - Section - \secref{sec:basics} explains how to use the framework naively, - presuming a reasonable default setup. Then, section - \secref{sec:advanced} deals with advanced topics, - introducing further aspects of the code generator framework - in a motivation-driven manner. Last, section \secref{sec:ml} - introduces the framework's internal programming interfaces. - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the already established code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. - Also note that while the framework itself is - object-logic independent, only \isa{HOL} provides a reasonable - framework setup. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{An example: a simple theory of search trees \label{sec:example}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When writing executable specifications using \isa{HOL}, - it is convenient to use - three existing packages: the datatype package for defining - datatypes, the function package for (recursive) functions, - and the class package for overloaded definitions. - - We develope a small theory of search trees; trees are represented - as a datatype with key type \isa{{\isacharprime}a} and value type \isa{{\isacharprime}b}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isacharequal}\ Leaf\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder{\isachardoublequoteclose}\ {\isacharprime}b\isanewline -\ \ {\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Note that we have constrained the type of keys - to the class of total orders, \isa{linorder}. - - We define \isa{find} and \isa{update} functions:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isasymColon}linorder{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}find\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isacharequal}\ key\ then\ Some\ val\ else\ None{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}find\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isasymle}\ key\ then\ find\ t{\isadigit{1}}\ it\ else\ find\ t{\isadigit{2}}\ it{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\isanewline -\ \ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline -\ \ \ \ if\ it\ {\isacharequal}\ key\ then\ Leaf\ key\ entry\isanewline -\ \ \ \ \ \ else\ if\ it\ {\isasymle}\ key\isanewline -\ \ \ \ \ \ then\ Branch\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\isanewline -\ \ \ \ \ \ else\ Branch\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\isanewline -\ \ \ {\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline -\ \ \ \ if\ it\ {\isasymle}\ key\isanewline -\ \ \ \ \ \ then\ {\isacharparenleft}Branch\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{1}}{\isacharparenright}\ key\ t{\isadigit{2}}{\isacharparenright}\isanewline -\ \ \ \ \ \ else\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{2}}{\isacharparenright}{\isacharparenright}\isanewline -\ \ \ {\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent For testing purpose, we define a small example - using natural numbers \isa{nat} (which are a \isa{linorder}) - as keys and list of nats as values:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ example\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat{\isacharcomma}\ nat\ list{\isacharparenright}\ searchtree{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}Leaf\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Then we generate code% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ example\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}tree{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent which looks like: - \lstsml{Thy/examples/tree.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Code generation concepts and process \label{sec:concept}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{figure}[h] - \centering - \includegraphics[width=0.7\textwidth]{codegen_process} - \caption{code generator -- processing overview} - \label{fig:process} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{defining equations}, \emph{datatypes}, and - \emph{type classes}. A defining equation as a first approximation - is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} - (an equation headed by a constant \isa{f} with arguments - \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). - Code generation aims to turn defining equations - into a functional program by running through - a process (see figure \ref{fig:process}): - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modeling - defining equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item These \qn{code theorems} then are \qn{translated} - into an Haskell-like intermediate - language. - - \item Finally, out of the intermediate language the final - code in the desired \qn{target language} is \qn{serialized}. - - \end{itemize} - - From these steps, only the two last are carried out - outside the logic; by keeping this layer as - thin as possible, the amount of code to trust is - kept to a minimum.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Basics \label{sec:basics}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Invoking the code generator% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Thanks to a reasonable setup of the \isa{HOL} theories, in - most cases code generation proceeds without further ado:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ fac\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}fac\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}fac\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharasterisk}\ fac\ n{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent This executable specification is now turned to SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ fac\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fac{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent The \isa{{\isasymEXPORTCODE}} command takes a space-separated list of - constants together with \qn{serialization directives} - These start with a \qn{target language} - identifier, followed by a file specification - where to write the generated code to. - - Internally, the defining equations for all selected - constants are taken, including any transitively required - constants, datatypes and classes, resulting in the following - code: - - \lstsml{Thy/examples/fac.ML} - - The code generator will complain when a required - ingredient does not provide a executable counterpart, - e.g.~generating code - for constants not yielding - a defining equation (e.g.~the Hilbert choice - operation \isa{SOME}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ pick{\isacharunderscore}some\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pick{\isacharunderscore}some\ xs\ {\isacharequal}\ {\isacharparenleft}SOME\ x{\isachardot}\ x\ {\isasymin}\ set\ xs{\isacharparenright}{\isachardoublequoteclose}% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ pick{\isacharunderscore}some\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fail{\isacharunderscore}const{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent will fail.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Theorem selection% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The list of all defining equations in a theory may be inspected - using the \isa{{\isasymPRINTCODESETUP}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{print{\isacharunderscore}codesetup}\isamarkupfalse% -% -\begin{isamarkuptext}% -\noindent which displays a table of constant with corresponding - defining equations (the additional stuff displayed - shall not bother us for the moment). - - The typical \isa{HOL} tools are already set up in a way that - function definitions introduced by \isa{{\isasymDEFINITION}}, - \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, - \isa{{\isasymRECDEF}} are implicitly propagated - to this defining equation table. Specific theorems may be - selected using an attribute: \emph{code func}. As example, - a weight selector function:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}let\ {\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}\ {\isacharequal}\ x\ in\isanewline -\ \ \ \ if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We want to eliminate the explicit destruction - of \isa{x} to \isa{{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ pick\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}pick{\isadigit{1}}{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent This theorem now is used for generating code: - - \lstsml{Thy/examples/pick1.ML} - - \noindent The policy is that \emph{default equations} stemming from - \isa{{\isasymDEFINITION}}, - \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, - \isa{{\isasymRECDEF}} statements are discarded as soon as an - equation is explicitly selected by means of \emph{code func}. - Further applications of \emph{code func} add theorems incrementally, - but syntactic redundancies are implicitly dropped. For example, - using a modified version of the \isa{fac} function - as defining equation, the then redundant (since - syntactically subsumed) original defining equations - are dropped. - - \begin{warn} - The attributes \emph{code} and \emph{code del} - associated with the existing code generator also apply to - the new one: \emph{code} implies \emph{code func}, - and \emph{code del} implies \emph{code func del}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Type classes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Type classes enter the game via the Isar class package. - For a short introduction how to use it, see \cite{isabelle-classes}; - here we just illustrate its impact on code generation. - - In a target language, type classes may be represented - natively (as in the case of Haskell). For languages - like SML, they are implemented using \emph{dictionaries}. - Our following example specifies a class \qt{null}, - assigning to each of its inhabitants a \qt{null} value:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{class}\isamarkupfalse% -\ null\ {\isacharequal}\ type\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ null\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}head\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ null{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}head\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We provide some instances for our \isa{null}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instantiation}\isamarkupfalse% -\ option\ \isakeyword{and}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ null\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ None{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -\noindent Constructing a dummy example:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}dummy\ {\isacharequal}\ head\ {\isacharbrackleft}Some\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ None{\isacharbrackright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Type classes offer a suitable occasion to introduce - the Haskell serializer. Its usage is almost the same - as SML, but, in accordance with conventions - some Haskell systems enforce, each module ends - up in a single file. The module hierarchy is reflected in - the file system, with root directory given as file specification.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ Haskell\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lsthaskell{Thy/examples/Codegen.hs} - \noindent (we have left out all other modules). - - \medskip - - The whole code in SML with explicit dictionary passing:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/class.ML} - - \medskip - - \noindent or in OCaml:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ OCaml\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ocaml{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/class.ocaml} - - \medskip The explicit association of constants - to classes can be inspected using the \isa{{\isasymPRINTCLASSES}} - command.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Recipes and advanced topics \label{sec:advanced}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In this tutorial, we do not attempt to give an exhaustive - description of the code generator framework; instead, - we cast a light on advanced topics by introducing - them together with practically motivated examples. Concerning - further reading, see - - \begin{itemize} - - \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} - for exhaustive syntax diagrams. - \item or \cite{Haftmann-Nipkow:2007:codegen} which deals with foundational issues - of the code generator framework. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Library theories \label{sec:library}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \isa{HOL} \isa{Main} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the \isa{HOL} - library; beside being useful in applications, they may serve - as a tutorial for customizing the code generator setup. - - \begin{description} - - \item[\isa{Code{\isacharunderscore}Integer}] represents \isa{HOL} integers by big - integer literals in target languages. - \item[\isa{Code{\isacharunderscore}Char}] represents \isa{HOL} characters by - character literals in target languages. - \item[\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}] like \isa{Code{\isacharunderscore}Char}, - but also offers treatment of character codes; includes - \isa{Code{\isacharunderscore}Integer}. - \item[\isa{Efficient{\isacharunderscore}Nat}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficency; pattern - matching with \isa{{\isadigit{0}}} / \isa{Suc} - is eliminated; includes \isa{Code{\isacharunderscore}Integer}. - \item[\isa{Code{\isacharunderscore}Index}] provides an additional datatype - \isa{index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[\isa{Code{\isacharunderscore}Message}] provides an additional datatype - \isa{message{\isacharunderscore}string} which is isomorphic to strings; - \isa{message{\isacharunderscore}string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Preprocessing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as defining equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared an undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - Some common applications:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{itemize} -% -\begin{isamarkuptext}% -\item replacing non-executable constructs by executable ones:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\item eliminating superfluous constants:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\item replacing executable but inconvenient constructs:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{itemize} -% -\begin{isamarkuptext}% -\emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} - pattern elimination implemented in - theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the \isa{{\isasymPRINTCODESETUP}} command. - - \begin{warn} - The attribute \emph{code unfold} - associated with the existing code generator also applies to - the new one: \emph{code unfold} implies \emph{code inline}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Concerning operational equality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Surely you have already noticed how equality is treated - by the code generator:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline -\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline -\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline -\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline -\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The membership test during preprocessing is rewritten, - resulting in \isa{op\ mem}, which itself - performs an explicit equality check.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ collect{\isacharunderscore}duplicates\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}collect{\isacharunderscore}duplicates{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/collect_duplicates.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class \isa{eq} with a corresponding operation - \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}}. - The preprocessing framework does the rest. - For datatypes, instances of \isa{eq} are implicitly derived - when possible. For other types, you may instantiate \isa{eq} - manually like any other type class. - - Though this \isa{eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ord{\isacharcomma}\ ord{\isacharparenright}\ ord\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ ord{\isacharunderscore}prod\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Then code generation will fail. Why? The definition - of \isa{op\ {\isasymle}} depends on equality on both arguments, - which are polymorphic and impose an additional \isa{eq} - class constraint, thus violating the type discipline - for class operations. - - The solution is to add \isa{eq} explicitly to the first sort arguments in the - code theorems:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ ord{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ ord{\isacharunderscore}prod\ \isacommand{by}\isamarkupfalse% -\ rule{\isacharplus}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Then code generation succeeds:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}ord\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}lexicographic{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/lexicographic.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -In general, code theorems for overloaded constants may have more - restrictive sort constraints than the underlying instance relation - between class and type constructor as long as the whole system of - constraints is coregular; code theorems violating coregularity - are rejected immediately. Consequently, it might be necessary - to delete disturbing theorems in the code theorem table, - as we have done here with the original definitions \isa{less{\isacharunderscore}prod{\isacharunderscore}def} - and \isa{less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def}. - - In some cases, the automatically derived defining equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers):% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Then code generation for SML would fail with a message - that the generated code conains illegal mutual dependencies: - the theorem \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}} already requires the - instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires - \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually - recursive \isa{instance} and \isa{function} definitions, - but the SML serializer does not support this. - - In such cases, you have to provide you own equality equations - involving auxiliary constants. In our case, - \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ {\isacharparenleft}op\ {\isacharequal}{\isacharparenright}\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}\ {\isacharcolon}{\isacharcolon}\ monotype\ {\isasymRightarrow}\ monotype\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}monotype{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/monotype.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Programs as sets of theorems% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the \isa{{\isasymCODETHMS}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ mod\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent prints a table with \emph{all} defining equations - for \isa{op\ mod}, including - \emph{all} defining equations those equations depend - on recursivly. \isa{{\isasymCODETHMS}} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on defining equations. - - Similarly, the \isa{{\isasymCODEDEPS}} command shows a graph - visualizing dependencies between defining equations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Constructor sets for datatypes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Conceptually, any datatype is spanned by a set of - \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} - where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is excactly the set of \emph{all} - type variables in \isa{{\isasymtau}}. The HOL datatype package - by default registers any new datatype in the table - of datatypes, which may be inspected using - the \isa{{\isasymPRINTCODESETUP}} command. - - In some cases, it may be convenient to alter or - extend this table; as an example, we will develope an alternative - representation of natural numbers as binary digits, whose - size does increase logarithmically with its value, not linear - \footnote{Indeed, the \isa{Efficient{\isacharunderscore}Nat} theory (see \ref{eff_nat}) - does something similar}. First, the digit representation:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\ Dig{\isadigit{0}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ n{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ Dig{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We will use these two ">digits"< to represent natural numbers - in binary digits, e.g.:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isadigit{4}}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Of course we also have to provide proper code equations for - the operations, e.g. \isa{op\ {\isacharplus}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}m\ {\isacharplus}\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{1}}\ m{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent We then instruct the code generator to view \isa{{\isadigit{0}}}, - \isa{{\isadigit{1}}}, \isa{Dig{\isadigit{0}}} and \isa{Dig{\isadigit{1}}} as - datatype constructors:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{0}}{\isasymColon}nat{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isadigit{1}}{\isasymColon}nat{\isachardoublequoteclose}\ Dig{\isadigit{0}}\ Dig{\isadigit{1}}% -\begin{isamarkuptext}% -\noindent For the former constructor \isa{Suc}, we provide a code - equation and remove some parts of the default code generator setup - which are an obstacle here:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ Suc{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}Suc\ n\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{declare}\isamarkupfalse% -\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline\ del{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}% -\begin{isamarkuptext}% -\noindent This yields the following code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharplus}\ {\isasymColon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}nat{\isacharunderscore}binary{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/nat_binary.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip - - From this example, it can be easily glimpsed that using own constructor sets - is a little delicate since it changes the set of valid patterns for values - of that type. Without going into much detail, here some practical hints: - - \begin{itemize} - \item When changing the constuctor set for datatypes, take care to - provide an alternative for the \isa{case} combinator (e.g. by replacing - it using the preprocessor). - \item Values in the target language need not to be normalized -- different - values in the target language may represent the same value in the - logic (e.g. \isa{Dig{\isadigit{1}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}}). - \item Usually, a good methodology to deal with the subleties of pattern - matching is to see the type as an abstract type: provide a set - of operations which operate on the concrete representation of the type, - and derive further operations by combinations of these primitive ones, - without relying on a particular representation. - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isachardoublequoteclose}\ Suc\isanewline -\isacommand{declare}\isamarkupfalse% -\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func{\isacharbrackright}\ \isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymColon}\ nat{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsubsection{Customizing serialization% -} -\isamarkuptrue% -% -\isamarkupsubsubsection{Basics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function and its corresponding - SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}literal{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_literal.ML} - - \noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialized as distinguished - entities with have nothing to do with the SML-builtin notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML \qt{bool} would be used. To map - the HOL \qt{bool} on SML \qt{bool}, we may use - \qn{custom serializations}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bool\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The \isa{{\isasymCODETYPE}} commad takes a type constructor - as arguments together with a list of custom serializations. - Each custom serialization starts with a target language - identifier followed by an expression, which during - code serialization is inserted whenever the type constructor - would occur. For constants, \isa{{\isasymCODECONST}} implements - the corresponding mechanism. Each ``\verb|_|'' in - a serialization expression is treated as a placeholder - for the type constructor's (the constant's) arguments.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}mlbool{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_mlbool.ML} - - \noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serializer - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -\isanewline -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}infix{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_infix.ML} - - \medskip - - Next, we try to map HOL pairs to SML pairs, using the - infix ``\verb|*|'' type constructor and parentheses:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ {\isacharasterisk}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ Pair\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The initial bang ``\verb|!|'' tells the serializer to never put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``\verb|/|'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serializations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serializations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialization may be quoted - using ``\verb|'|''; thus, in - ``\verb|fn '_ => _|'' the first - ``\verb|_|'' is a proper underscore while the - second ``\verb|_|'' is a placeholder. - - The HOL theories provide further - examples for custom serializations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Haskell serialization% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For convenience, the default - HOL setup for Haskell maps the \isa{eq} class to - its counterpart in Haskell, giving custom serializations - for the class (\isa{{\isasymCODECLASS}}) and its operation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}class}\isamarkupfalse% -\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}\ \isakeyword{where}\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharequal}{\isacharequal}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -A problem now occurs whenever a type which - is an instance of \isa{eq} in HOL is mapped - on a Haskell-builtin type which is also an instance - of Haskell \isa{Eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\isamarkupfalse% -\ bar\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -% -\isadelimtt -\isanewline -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bar\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The code generator would produce - an additional instance, which of course is rejected. - To suppress this additional instance, use - \isa{{\isasymCODEINSTANCE}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\isamarkupsubsubsection{Pretty printing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The serializer provides ML interfaces to set up - pretty serializations for expressions like lists, numerals - and characters; these are - monolithic stubs and should only be used with the - theories introduced in \secref{sec:library}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Cyclic module dependencies% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Sometimes the awkward situation occurs that dependencies - between definitions introduce cyclic dependencies - between modules, which in the Haskell world leaves - you to the mercy of the Haskell implementation you are using, - while for SML code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% -\ SML\isanewline -\ \ A\ ABC\isanewline -\ \ B\ ABC\isanewline -\ \ C\ ABC% -\begin{isamarkuptext}% -we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialization time.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Incremental code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Code generation is \emph{incremental}: theorems - and abstract intermediate code are cached and extended on demand. - The cache may be partially or fully dropped if the underlying - executable content of the theory changes. - Implementation of caching is supposed to transparently - hid away the details from the user. Anyway, caching - reaches the surface by using a slightly more general form - of the \isa{{\isasymCODETHMS}}, \isa{{\isasymCODEDEPS}} - and \isa{{\isasymEXPORTCODE}} commands: the list of constants - may be omitted. Then, all constants with code theorems - in the current cache are referred to.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{ML interfaces \label{sec:ml}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Executable theory content: \isa{Code}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This Pure module implements the core notions of - executable content of a theory.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Managing executable content% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code.add\_func}\verb|Code.add_func: thm -> theory -> theory| \\ - \indexml{Code.del\_func}\verb|Code.del_func: thm -> theory -> theory| \\ - \indexml{Code.add\_funcl}\verb|Code.add_funcl: string * thm list Susp.T -> theory -> theory| \\ - \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ - \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ - \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> thm list -> thm list option)|\isasep\isanewline% -\verb| -> theory -> theory| \\ - \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ - \indexml{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ - \indexml{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% -\verb| -> (string * sort) list * (string * typ list) list| \\ - \indexml{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| - \end{mldecls} - - \begin{description} - - \item \verb|Code.add_func|~\isa{thm}~\isa{thy} adds function - theorem \isa{thm} to executable content. - - \item \verb|Code.del_func|~\isa{thm}~\isa{thy} removes function - theorem \isa{thm} from executable content, if present. - - \item \verb|Code.add_funcl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds - suspended defining equations \isa{lthms} for constant - \isa{const} to executable content. - - \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes - the preprocessor simpset. - - \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds - function transformer \isa{f} (named \isa{name}) to executable content; - \isa{f} is a transformer of the defining equations belonging - to a certain function definition, depending on the - current theory context. Returning \isa{NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new defining equations. - - \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes - function transformer named \isa{name} from executable content. - - \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds - a datatype to executable content, with generation - set \isa{cs}. - - \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} - returns type constructor corresponding to - constructor \isa{const}; returns \isa{NONE} - if \isa{const} is no constructor. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Auxiliary% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ - \indexml{Code\_Unit.head\_func}\verb|Code_Unit.head_func: thm -> string * ((string * sort) list * typ)| \\ - \indexml{Code\_Unit.rewrite\_func}\verb|Code_Unit.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} - reads a constant as a concrete term expression \isa{s}. - - \item \verb|Code_Unit.head_func|~\isa{thm} - extracts the constant and its type from a defining equation \isa{thm}. - - \item \verb|Code_Unit.rewrite_func|~\isa{ss}~\isa{thm} - rewrites a defining equation \isa{thm} with a simpset \isa{ss}; - only arguments and right hand side are rewritten, - not the head of the defining equation. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Implementing code generator applications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things. - - \begin{warn} - Some interfaces discussed here have not reached - a final state yet. - Changes likely to occur in future. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Data depending on the theory's executable content% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot \verb|CodeDataFun|; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - \isa{type\ T} \\ - \isa{val\ empty{\isacharcolon}\ T} \\ - \isa{val\ merge{\isacharcolon}\ Pretty{\isachardot}pp\ {\isasymrightarrow}\ T\ {\isacharasterisk}\ T\ {\isasymrightarrow}\ T} \\ - \isa{val\ purge{\isacharcolon}\ theory\ option\ {\isasymrightarrow}\ CodeUnit{\isachardot}const\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} - \end{tabular} - - \begin{description} - - \item \isa{T} the type of data to store. - - \item \isa{empty} initial (empty) data. - - \item \isa{merge} merging two data slots. - - \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; - if possible, the current theory context is handed over - as argument \isa{thy} (if there is no current theory context (e.g.~during - theory merge, \verb|NONE|); \isa{consts} indicates the kind - of change: \verb|NONE| stands for a fundamental change - which invalidates any existing code, \isa{SOME\ consts} - hints that executable content for constants \isa{consts} - has changed. - - \end{description} - - An instance of \verb|CodeDataFun| provides the following - interface: - - \medskip - \begin{tabular}{l} - \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} - \end{tabular} - - \begin{description} - - \item \isa{get} retrieval of the current data. - - \item \isa{change} update of current data (cached!) - by giving a continuation. - - \item \isa{change{\isacharunderscore}yield} update with side result. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\emph{Happy proving, happy hacking!}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/document/Further.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Further.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,234 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Further}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Further\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{Further issues \label{sec:further}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Further reading% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Do dive deeper into the issue of code generation, you should visit - the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which - contains exhaustive syntax diagrams.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Modules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When invoking the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command it is possible to leave - out the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isacharunderscore}name}}}} part; then code is distributed over - different modules, where the module name space roughly is induced - by the \isa{Isabelle} theory name space. - - Then sometimes the awkward situation occurs that dependencies between - definitions introduce cyclic dependencies between modules, which in the - \isa{Haskell} world leaves you to the mercy of the \isa{Haskell} implementation - you are using, while for \isa{SML}/\isa{OCaml} code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% -\ SML\isanewline -\ \ A\ ABC\isanewline -\ \ B\ ABC\isanewline -\ \ C\ ABC% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialisation time.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Evaluation oracle% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Code generation may also be used to \emph{evaluate} expressions - (using \isa{SML} as target language of course). - For instance, the \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} allows to reduce an expression to a - normal form with respect to the underlying code equations:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{value}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent will display \isa{{\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}}. - - The \hyperlink{method.eval}{\mbox{\isa{eval}}} method tries to reduce a goal by code generation to \isa{True} - and solves it in that case, but fails otherwise:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}\ {\isacharequal}\ {\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ eval% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The soundness of the \hyperlink{method.eval}{\mbox{\isa{eval}}} method depends crucially - on the correctness of the code generator; this is one of the reasons - why you should not use adaption (see \secref{sec:adaption}) frivolously.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Code antiquotation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In scenarios involving techniques like reflection it is quite common - that code generated from a theory forms the basis for implementing - a proof procedure in \isa{SML}. To facilitate interfacing of generated code - with system code, the code generator provides a \isa{code} antiquotation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ form\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ And\ form\ form\ {\isacharbar}\ Or\ form\ form\isanewline -\isanewline -\isacommand{ML}\isamarkupfalse% -\ {\isacharverbatimopen}\isanewline -\ \ fun\ eval{\isacharunderscore}form\ % -\isaantiq -code\ T% -\endisaantiq -\ {\isacharequal}\ true\isanewline -\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ % -\isaantiq -code\ F% -\endisaantiq -\ {\isacharequal}\ false\isanewline -\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% -\isaantiq -code\ And% -\endisaantiq -\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ andalso\ eval{\isacharunderscore}form\ q\isanewline -\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% -\isaantiq -code\ Or% -\endisaantiq -\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ orelse\ eval{\isacharunderscore}form\ q{\isacharsemicolon}\isanewline -{\isacharverbatimclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent \isa{code} takes as argument the name of a constant; after the - whole \isa{SML} is read, the necessary code is generated transparently - and the corresponding constant names are inserted. This technique also - allows to use pattern matching on constructors stemming from compiled - \isa{datatypes}. - - For a less simplistic example, theory \hyperlink{theory.Ferrack}{\mbox{\isa{Ferrack}}} is - a good reference.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Imperative data structures% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -If you consider imperative data structures as inevitable for a specific - application, you should consider - \emph{Imperative Functional Programming with Isabelle/HOL} - (\cite{bulwahn-et-al:2008:imperative}); - the framework described there is available in theory \hyperlink{theory.Imperative-HOL}{\mbox{\isa{Imperative{\isacharunderscore}HOL}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,390 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Introduction}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Introduction\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Code generation from \isa{Isabelle{\isacharslash}HOL} theories% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction and Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This tutorial introduces a generic code generator for the - \isa{Isabelle} system. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and \isa{Haskell} - \cite{haskell-revised-report}). - - Conceptually the code generator framework is part - of Isabelle's \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} meta logic framework; the logic - \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} which is an extension of \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} distribution theories. - (see also \cite{isa-tutorial}). - - The code generator aims to be usable with no further ado - in most cases while allowing for detailed customisation. - This manifests in the structure of this tutorial: after a short - conceptual introduction with an example (\secref{sec:intro}), - we discuss the generic customisation facilities (\secref{sec:program}). - A further section (\secref{sec:adaption}) is dedicated to the matter of - \qn{adaption} to specific target language environments. After some - further issues (\secref{sec:further}) we conclude with an overview - of some ML programming interfaces (\secref{sec:ml}). - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the existing code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. In case of ambiguity, we will refer to the framework - described here as \isa{generic\ code\ generator}, to the - other as \isa{SML\ code\ generator}. - Also note that while the framework itself is - object-logic independent, only \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} provides a reasonable - framework setup. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Code generation via shallow embedding \label{sec:intro}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The key concept for understanding \isa{Isabelle}'s code generation is - \emph{shallow embedding}, i.e.~logical entities like constants, types and - classes are identified with corresponding concepts in the target language. - - Inside \hyperlink{theory.HOL}{\mbox{\isa{HOL}}}, the \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} and - \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} declarations form - the core of a functional programming language. The default code generator setup - allows to turn those into functional programs immediately. - This means that \qt{naive} code generation can proceed without further ado. - For example, here a simple \qt{implementation} of amortised queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ {\isacharprime}a\ queue\ {\isacharequal}\ AQueue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then we can generate code e.g.~for \isa{SML} as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ empty\ dequeue\ enqueue\ \isakeyword{in}\ SML\isanewline -\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}example{\isachardot}ML{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent resulting in the following code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun foldl f a [] = a\\ -\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun list{\char95}case f1 f2 (a ::~lista) = f2 a lista\\ -\hspace*{0pt} ~| list{\char95}case f1 f2 [] = f1;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun dequeue (AQueue ([],~[])) = (NONE,~AQueue ([],~[]))\\ -\hspace*{0pt} ~| dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ -\hspace*{0pt} ~| dequeue (AQueue (v ::~va,~[])) =\\ -\hspace*{0pt} ~~~let\\ -\hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\ -\hspace*{0pt} ~~~in\\ -\hspace*{0pt} ~~~~~(SOME y,~AQueue ([],~ys))\\ -\hspace*{0pt} ~~~end;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command takes a space-separated list of - constants for which code shall be generated; anything else needed for those - is added implicitly. Then follows a target language identifier - (\isa{SML}, \isa{OCaml} or \isa{Haskell}) and a freely chosen module name. - A file name denotes the destination to store the generated code. Note that - the semantics of the destination depends on the target language: for - \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell} - it denotes a \emph{directory} where a file named as the module name - (with extension \isa{{\isachardot}hs}) is written:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ empty\ dequeue\ enqueue\ \isakeyword{in}\ Haskell\isanewline -\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This is how the corresponding code in \isa{Haskell} looks like:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}module Example where {\char123}\\ -\hspace*{0pt}\\ -\hspace*{0pt}\\ -\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\ -\hspace*{0pt}foldla f a [] = a;\\ -\hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\ -\hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}list{\char95}case ::~forall t a.~t -> (a -> [a] -> t) -> [a] -> t;\\ -\hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\ -\hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\ -\hspace*{0pt}\\ -\hspace*{0pt}data Queue a = AQueue [a] [a];\\ -\hspace*{0pt}\\ -\hspace*{0pt}empty ::~forall a.~Queue a;\\ -\hspace*{0pt}empty = AQueue [] [];\\ -\hspace*{0pt}\\ -\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ -\hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\ -\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ -\hspace*{0pt}dequeue (AQueue (v :~va) []) =\\ -\hspace*{0pt} ~let {\char123}\\ -\hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\ -\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\ -\hspace*{0pt}\\ -\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\ -\hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\ -\hspace*{0pt}\\ -\hspace*{0pt}{\char125}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This demonstrates the basic usage of the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command; - for more details see \secref{sec:further}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Code generator architecture \label{sec:concept}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -What you have seen so far should be already enough in a lot of cases. If you - are content with this, you can quit reading here. Anyway, in order to customise - and adapt the code generator, it is inevitable to gain some understanding - how it works. - - \begin{figure}[h] - \begin{tikzpicture}[x = 4.2cm, y = 1cm] - \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; - \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; - \tikzstyle process_arrow=[->, semithick, color = green]; - \node (HOL) at (0, 4) [style=entity] {\isa{Isabelle{\isacharslash}HOL} theory}; - \node (eqn) at (2, 2) [style=entity] {code equations}; - \node (iml) at (2, 0) [style=entity] {intermediate language}; - \node (seri) at (1, 0) [style=process] {serialisation}; - \node (SML) at (0, 3) [style=entity] {\isa{SML}}; - \node (OCaml) at (0, 2) [style=entity] {\isa{OCaml}}; - \node (further) at (0, 1) [style=entity] {\isa{{\isasymdots}}}; - \node (Haskell) at (0, 0) [style=entity] {\isa{Haskell}}; - \draw [style=process_arrow] (HOL) .. controls (2, 4) .. - node [style=process, near start] {selection} - node [style=process, near end] {preprocessing} - (eqn); - \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); - \draw [style=process_arrow] (iml) -- (seri); - \draw [style=process_arrow] (seri) -- (SML); - \draw [style=process_arrow] (seri) -- (OCaml); - \draw [style=process_arrow, dashed] (seri) -- (further); - \draw [style=process_arrow] (seri) -- (Haskell); - \end{tikzpicture} - \caption{Code generator architecture} - \label{fig:arch} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{code equations}, \emph{datatypes}, and - \emph{type classes}. A code equation as a first approximation - is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} - (an equation headed by a constant \isa{f} with arguments - \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). - Code generation aims to turn code equations - into a functional program. This is achieved by three major - components which operate sequentially, i.e. the result of one is - the input - of the next in the chain, see diagram \ref{fig:arch}: - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modelling - code equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item Before the selected code equations are continued with, - they can be \qn{preprocessed}, i.e. subjected to theorem - transformations. This \qn{preprocessor} is an interface which - allows to apply - the full expressiveness of ML-based theorem transformations - to code generation; motivating examples are shown below, see - \secref{sec:preproc}. - The result of the preprocessing step is a structured collection - of code equations. - - \item These code equations are \qn{translated} to a program - in an abstract intermediate language. Think of it as a kind - of \qt{Mini-Haskell} with four \qn{statements}: \isa{data} - (for datatypes), \isa{fun} (stemming from code equations), - also \isa{class} and \isa{inst} (for type classes). - - \item Finally, the abstract program is \qn{serialised} into concrete - source code of a target language. - - \end{itemize} - - \noindent From these steps, only the two last are carried out outside the logic; by - keeping this layer as thin as possible, the amount of code to trust is - kept to a minimum.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,255 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{ML}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ {\isachardoublequoteopen}ML{\isachardoublequoteclose}\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{ML system interfaces \label{sec:ml}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Executable theory content: \isa{Code}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This Pure module implements the core notions of - executable content of a theory.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Managing executable content% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\ - \indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\ - \indexdef{}{ML}{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\ - \indexdef{}{ML}{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\ - \indexdef{}{ML}{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\ - \indexdef{}{ML}{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline% -\verb| -> theory -> theory| \\ - \indexdef{}{ML}{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ - \indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ - \indexdef{}{ML}{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% -\verb| -> (string * sort) list * (string * typ list) list| \\ - \indexdef{}{ML}{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| - \end{mldecls} - - \begin{description} - - \item \verb|Code.add_eqn|~\isa{thm}~\isa{thy} adds function - theorem \isa{thm} to executable content. - - \item \verb|Code.del_eqn|~\isa{thm}~\isa{thy} removes function - theorem \isa{thm} from executable content, if present. - - \item \verb|Code.add_eqnl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds - suspended code equations \isa{lthms} for constant - \isa{const} to executable content. - - \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes - the preprocessor simpset. - - \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds - function transformer \isa{f} (named \isa{name}) to executable content; - \isa{f} is a transformer of the code equations belonging - to a certain function definition, depending on the - current theory context. Returning \isa{NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new code equations. - - \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes - function transformer named \isa{name} from executable content. - - \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds - a datatype to executable content, with generation - set \isa{cs}. - - \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} - returns type constructor corresponding to - constructor \isa{const}; returns \isa{NONE} - if \isa{const} is no constructor. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Auxiliary% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ - \indexdef{}{ML}{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\ - \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} - reads a constant as a concrete term expression \isa{s}. - - \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm} - extracts the constant and its type from a code equation \isa{thm}. - - \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm} - rewrites a code equation \isa{thm} with a simpset \isa{ss}; - only arguments and right hand side are rewritten, - not the head of the code equation. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Implementing code generator applications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Data depending on the theory's executable content% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot \verb|CodeDataFun|; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - \isa{type\ T} \\ - \isa{val\ empty{\isacharcolon}\ T} \\ - \isa{val\ purge{\isacharcolon}\ theory\ {\isasymrightarrow}\ string\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} - \end{tabular} - - \begin{description} - - \item \isa{T} the type of data to store. - - \item \isa{empty} initial (empty) data. - - \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; - \isa{consts} indicates the kind - of change: \verb|NONE| stands for a fundamental change - which invalidates any existing code, \isa{SOME\ consts} - hints that executable content for constants \isa{consts} - has changed. - - \end{description} - - \noindent An instance of \verb|CodeDataFun| provides the following - interface: - - \medskip - \begin{tabular}{l} - \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} - \end{tabular} - - \begin{description} - - \item \isa{get} retrieval of the current data. - - \item \isa{change} update of current data (cached!) - by giving a continuation. - - \item \isa{change{\isacharunderscore}yield} update with side result. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\bigskip - - \emph{Happy proving, happy hacking!}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1250 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Program}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Program\isanewline -\isakeyword{imports}\ Introduction\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{Turning Theories into Programs \label{sec:program}% -} -\isamarkuptrue% -% -\isamarkupsubsection{The \isa{Isabelle{\isacharslash}HOL} default setup% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We have already seen how by default equations stemming from - \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} - statements are used for code generation. This default behaviour - can be changed, e.g. by providing different code equations. - All kinds of customisation shown in this section is \emph{safe} - in the sense that the user does not have to worry about - correctness -- all programs generatable that way are partially - correct.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Selecting code equations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Coming back to our introductory example, we - could provide an alternative code equations for \isa{dequeue} - explicitly:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}cases\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\ {\isacharparenleft}cases\ {\isachardoublequoteopen}rev\ xs{\isachardoublequoteclose}{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The annotation \isa{{\isacharbrackleft}code{\isacharbrackright}} is an \isa{Isar} - \isa{attribute} which states that the given theorems should be - considered as code equations for a \isa{fun} statement -- - the corresponding constant is determined syntactically. The resulting code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ -\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ -\hspace*{0pt}dequeue (AQueue xs []) =\\ -\hspace*{0pt} ~(if nulla xs then (Nothing,~AQueue [] [])\\ -\hspace*{0pt} ~~~else dequeue (AQueue [] (rev xs)));% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent You may note that the equality test \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} has been - replaced by the predicate \isa{null\ xs}. This is due to the default - setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). - - Changing the default constructor set of datatypes is also - possible. See \secref{sec:datatypes} for an example. - - As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% -\ dequeue% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent prints a table with \emph{all} code equations - for \isa{dequeue}, including - \emph{all} code equations those equations depend - on recursively. - - Similarly, the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command shows a graph - visualising dependencies between code equations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{\isa{class} and \isa{instantiation}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Concerning type classes and code generation, let us examine an example - from abstract algebra:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ semigroup\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{class}\isamarkupfalse% -\ monoid\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ nat\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isasymotimes}\ n\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ n\ {\isacharplus}\ m\ {\isasymotimes}\ n{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ neutral{\isacharunderscore}nat\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharcolon}\isanewline -\ \ \isakeyword{fixes}\ n\ m\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ m{\isacharparenright}\ {\isasymotimes}\ q\ {\isacharequal}\ n\ {\isasymotimes}\ q\ {\isacharplus}\ m\ {\isasymotimes}\ q{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}m\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent We define the natural operation of the natural numbers - on monoids:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}pow\ {\isadigit{0}}\ a\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}pow\ {\isacharparenleft}Suc\ n{\isacharparenright}\ a\ {\isacharequal}\ a\ {\isasymotimes}\ pow\ n\ a{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This we use to define the discrete exponentiation function:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ bexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}bexp\ n\ {\isacharequal}\ pow\ n\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The corresponding code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}module Example where {\char123}\\ -\hspace*{0pt}\\ -\hspace*{0pt}\\ -\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}class Semigroup a where {\char123}\\ -\hspace*{0pt} ~mult ::~a -> a -> a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Semigroup a) => Monoid a where {\char123}\\ -\hspace*{0pt} ~neutral ::~a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\ -\hspace*{0pt}pow Zero{\char95}nat a = neutral;\\ -\hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\ -\hspace*{0pt}\\ -\hspace*{0pt}plus{\char95}nat ::~Nat -> Nat -> Nat;\\ -\hspace*{0pt}plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n);\\ -\hspace*{0pt}plus{\char95}nat Zero{\char95}nat n = n;\\ -\hspace*{0pt}\\ -\hspace*{0pt}neutral{\char95}nat ::~Nat;\\ -\hspace*{0pt}neutral{\char95}nat = Suc Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}mult{\char95}nat ::~Nat -> Nat -> Nat;\\ -\hspace*{0pt}mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat;\\ -\hspace*{0pt}mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Semigroup Nat where {\char123}\\ -\hspace*{0pt} ~mult = mult{\char95}nat;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Monoid Nat where {\char123}\\ -\hspace*{0pt} ~neutral = neutral{\char95}nat;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}bexp ::~Nat -> Nat;\\ -\hspace*{0pt}bexp n = pow n (Suc (Suc Zero{\char95}nat));\\ -\hspace*{0pt}\\ -\hspace*{0pt}{\char125}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This is a convenient place to show how explicit dictionary construction - manifests in generated code (here, the same example in \isa{SML}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ -\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\ -\hspace*{0pt}fun semigroup{\char95}monoid (A{\char95}:'a monoid) = {\char35}Program{\char95}{\char95}semigroup{\char95}monoid A{\char95};\\ -\hspace*{0pt}fun neutral (A{\char95}:'a monoid) = {\char35}neutral A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun pow A{\char95}~Zero{\char95}nat a = neutral A{\char95}\\ -\hspace*{0pt} ~| pow A{\char95}~(Suc n) a = mult (semigroup{\char95}monoid A{\char95}) a (pow A{\char95}~n a);\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\ -\hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\ -\hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ -\hspace*{0pt}\\ -\hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoid{\char95}nat =\\ -\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}~:\\ -\hspace*{0pt} ~nat monoid;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Note the parameters with trailing underscore (\verb|A_|) - which are the dictionary parameters.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The preprocessor \label{sec:preproc}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as code equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared and undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - - Some common applications:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{itemize} -% -\begin{isamarkuptext}% -\item replacing non-executable constructs by executable ones:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\item eliminating superfluous constants:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\item replacing executable but inconvenient constructs:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\end{itemize} -% -\begin{isamarkuptext}% -\noindent \emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} - pattern elimination implemented in - theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. - \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on code equations. - - \begin{warn} - The attribute \emph{code unfold} - associated with the \isa{SML\ code\ generator} also applies to - the \isa{generic\ code\ generator}: - \emph{code unfold} implies \emph{code inline}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Datatypes \label{sec:datatypes}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Conceptually, any datatype is spanned by a set of - \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is exactly the set of \emph{all} type variables in - \isa{{\isasymtau}}. The HOL datatype package by default registers any new - datatype in the table of datatypes, which may be inspected using the - \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. - - In some cases, it is appropriate to alter or extend this table. As - an example, we will develop an alternative representation of the - queue example given in \secref{sec:intro}. The amortised - representation is convenient for generating code but exposes its - \qt{implementation} details, which may be cumbersome when proving - theorems about it. Therefore, here a simple, straightforward - representation of queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ {\isacharprime}a\ queue\ {\isacharequal}\ Queue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}Queue\ xs{\isacharparenright}\ {\isacharequal}\ Queue\ {\isacharparenleft}xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ x{\isacharcomma}\ Queue\ xs{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This we can use directly for proving; for executing, - we provide an alternative characterisation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ AQueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}AQueue\ xs\ ys\ {\isacharequal}\ Queue\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ AQueue% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Here we define a \qt{constructor} \isa{Program{\isachardot}AQueue} which - is defined in terms of \isa{Queue} and interprets its arguments - according to what the \emph{content} of an amortised queue is supposed - to be. Equipped with this, we are able to prove the following equations - for our primitive queue operations which \qt{implement} the simple - queues in an amortised fashion:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ empty{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ AQueue{\isacharunderscore}def\ empty{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ enqueue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent For completeness, we provide a substitute for the - \isa{case} combinator on queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ aqueue{\isacharunderscore}case{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}aqueue{\isacharunderscore}case\ {\isacharequal}\ queue{\isacharunderscore}case{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ aqueue{\isacharunderscore}case\ {\isacharbrackleft}code{\isacharcomma}\ code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}queue{\isacharunderscore}case\ {\isacharequal}\ aqueue{\isacharunderscore}case{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ aqueue{\isacharunderscore}case{\isacharunderscore}def\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ case{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}aqueue{\isacharunderscore}case\ f\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ aqueue{\isacharunderscore}case{\isacharunderscore}def\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The resulting code looks as expected:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun foldl f a [] = a\\ -\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun null [] = true\\ -\hspace*{0pt} ~| null (x ::~xs) = false;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ -\hspace*{0pt} ~| dequeue (AQueue (xs,~[])) =\\ -\hspace*{0pt} ~~~(if null xs then (NONE,~AQueue ([],~[]))\\ -\hspace*{0pt} ~~~~~else dequeue (AQueue ([],~rev xs)));\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent From this example, it can be glimpsed that using own - constructor sets is a little delicate since it changes the set of - valid patterns for values of that type. Without going into much - detail, here some practical hints: - - \begin{itemize} - - \item When changing the constructor set for datatypes, take care - to provide an alternative for the \isa{case} combinator - (e.g.~by replacing it using the preprocessor). - - \item Values in the target language need not to be normalised -- - different values in the target language may represent the same - value in the logic. - - \item Usually, a good methodology to deal with the subtleties of - pattern matching is to see the type as an abstract type: provide - a set of operations which operate on the concrete representation - of the type, and derive further operations by combinations of - these primitive ones, without relying on a particular - representation. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Equality and wellsortedness% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Surely you have already noticed how equality is treated - by the code generator:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline -\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline -\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline -\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline -\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The membership test during preprocessing is rewritten, - resulting in \isa{op\ mem}, which itself - performs an explicit equality check.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun member A{\char95}~x [] = false\\ -\hspace*{0pt} ~| member A{\char95}~x (y ::~ys) = eqop A{\char95}~x y orelse member A{\char95}~x ys;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun collect{\char95}duplicates A{\char95}~xs ys [] = xs\\ -\hspace*{0pt} ~| collect{\char95}duplicates A{\char95}~xs ys (z ::~zs) =\\ -\hspace*{0pt} ~~~(if member A{\char95}~z xs\\ -\hspace*{0pt} ~~~~~then (if member A{\char95}~z ys then collect{\char95}duplicates A{\char95}~xs ys zs\\ -\hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\ -\hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class \isa{eq} with a corresponding operation - \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}. - The preprocessing framework does the rest by propagating the - \isa{eq} constraints through all dependent code equations. - For datatypes, instances of \isa{eq} are implicitly derived - when possible. For other types, you may instantiate \isa{eq} - manually like any other type class. - - Though this \isa{eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples - (also see theory \hyperlink{theory.Product-ord}{\mbox{\isa{Product{\isacharunderscore}ord}}}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}order{\isacharcomma}\ order{\isacharparenright}\ order\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isasymle}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isasymle}\ snd\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isacharless}\ snd\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}auto\ simp{\isacharcolon}\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}prod{\isacharunderscore}def\ intro{\isacharcolon}\ order{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ order{\isacharunderscore}prod\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then code generation will fail. Why? The definition - of \isa{op\ {\isasymle}} depends on equality on both arguments, - which are polymorphic and impose an additional \isa{eq} - class constraint, which the preprocessor does not propagate - (for technical reasons). - - The solution is to add \isa{eq} explicitly to the first sort arguments in the - code theorems:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ order{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then code generation succeeds:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\ -\hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a preorder = {\char123}Orderings{\char95}{\char95}ord{\char95}preorder :~'a ord{\char125};\\ -\hspace*{0pt}fun ord{\char95}preorder (A{\char95}:'a preorder) = {\char35}Orderings{\char95}{\char95}ord{\char95}preorder A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\ -\hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ -\hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ -\hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\ -\hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\ -\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ -\hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ -\hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\ -\hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -In some cases, the automatically derived code equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Then code generation for SML would fail with a message - that the generated code contains illegal mutual dependencies: - the theorem \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}} already requires the - instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires - \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually - recursive \isa{instance} and \isa{function} definitions, - but the SML serialiser does not support this. - - In such cases, you have to provide your own equality equations - involving auxiliary constants. In our case, - \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun null [] = true\\ -\hspace*{0pt} ~| null (x ::~xs) = false;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eq{\char95}nat (Suc a) Zero{\char95}nat = false\\ -\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc a) = false\\ -\hspace*{0pt} ~| eq{\char95}nat (Suc nat) (Suc nat') = eq{\char95}nat nat nat'\\ -\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat Zero{\char95}nat = true;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype monotype = Mono of nat * monotype list;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun list{\char95}all2 p (x ::~xs) (y ::~ys) = p x y andalso list{\char95}all2 p xs ys\\ -\hspace*{0pt} ~| list{\char95}all2 p xs [] = null xs\\ -\hspace*{0pt} ~| list{\char95}all2 p [] ys = null ys;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1,~typargs1)) (Mono (tyco2,~typargs2)) =\\ -\hspace*{0pt} ~eq{\char95}nat tyco1 tyco2 andalso list{\char95}all2 eq{\char95}monotype typargs1 typargs2;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsubsection{Explicit partiality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Partiality usually enters the game by partial patterns, as - in the following example, again for amortised queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\isanewline -\ \ \ \ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent In the corresponding code, there is no equation - for the pattern \isa{Program{\isachardot}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\ -\hspace*{0pt} ~let {\char123}\\ -\hspace*{0pt} ~~~(y :~ys) = rev xs;\\ -\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent In some cases it is desirable to have this - pseudo-\qt{partiality} more explicitly, e.g.~as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{axiomatization}\isamarkupfalse% -\ empty{\isacharunderscore}queue\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isacharbar}\ {\isacharunderscore}\ {\isasymRightarrow}\ empty{\isacharunderscore}queue{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ empty{\isacharunderscore}queue\isanewline -\ \ \ \ \ else\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -Observe that on the right hand side of the definition of \isa{strict{\isacharunderscore}dequeue{\isacharprime}} the constant \isa{empty{\isacharunderscore}queue} occurs - which is unspecified. - - Normally, if constants without any code equations occur in a - program, the code generator complains (since in most cases this is - not what the user expects). But such constants can also be thought - of as function definitions with no equations which always fail, - since there is never a successful pattern match on the left hand - side. In order to categorise a constant into that category - explicitly, use \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}abort}\isamarkupfalse% -\ empty{\isacharunderscore}queue% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then the code generator will just insert an error or - exception at the appropriate position:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\ -\hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\ -\hspace*{0pt}\\ -\hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\ -\hspace*{0pt}strict{\char95}dequeue' (AQueue xs (y :~ys)) = (y,~AQueue xs ys);\\ -\hspace*{0pt}strict{\char95}dequeue' (AQueue xs []) =\\ -\hspace*{0pt} ~(if nulla xs then empty{\char95}queue\\ -\hspace*{0pt} ~~~else strict{\char95}dequeue' (AQueue [] (rev xs)));% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This feature however is rarely needed in practice. - Note also that the \isa{HOL} default setup already declares - \isa{undefined} as \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}, which is most - likely to be used in such situations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\ \end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/Codegen.hs --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/Codegen.hs Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,23 +0,0 @@ -module Codegen where { - -import qualified Nat; - -class Null a where { - nulla :: a; -}; - -heada :: forall a. (Codegen.Null a) => [a] -> a; -heada (x : xs) = x; -heada [] = Codegen.nulla; - -null_option :: forall a. Maybe a; -null_option = Nothing; - -instance Codegen.Null (Maybe a) where { - nulla = Codegen.null_option; -}; - -dummy :: Maybe Nat.Nat; -dummy = Codegen.heada [Just (Nat.Suc Nat.Zero_nat), Nothing]; - -} diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ -{-# OPTIONS_GHC -fglasgow-exts #-} - -module Example where { - - -foldla :: forall a b. (a -> b -> a) -> a -> [b] -> a; -foldla f a [] = a; -foldla f a (x : xs) = foldla f (f a x) xs; - -rev :: forall a. [a] -> [a]; -rev xs = foldla (\ xsa x -> x : xsa) [] xs; - -list_case :: forall t a. t -> (a -> [a] -> t) -> [a] -> t; -list_case f1 f2 (a : list) = f2 a list; -list_case f1 f2 [] = f1; - -data Queue a = AQueue [a] [a]; - -empty :: forall a. Queue a; -empty = AQueue [] []; - -dequeue :: forall a. Queue a -> (Maybe a, Queue a); -dequeue (AQueue [] []) = (Nothing, AQueue [] []); -dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys); -dequeue (AQueue (v : va) []) = - let { - (y : ys) = rev (v : va); - } in (Just y, AQueue [] ys); - -enqueue :: forall a. a -> Queue a -> Queue a; -enqueue x (AQueue xs ys) = AQueue (x : xs) ys; - -} diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/arbitrary.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/arbitrary.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,9 +0,0 @@ -structure Codegen = -struct - -val arbitrary_option : 'a option = NONE; - -fun dummy_option [] = arbitrary_option - | dummy_option (x :: xs) = SOME x; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/bool_infix.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/bool_infix.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun in_interval (k, l) n = - Nat.less_eq_nat k n andalso Nat.less_eq_nat n l; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/bool_literal.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/bool_literal.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,31 +0,0 @@ -structure HOL = -struct - -datatype boola = False | True; - -fun anda x True = x - | anda x False = False - | anda True x = x - | anda False x = False; - -end; (*struct HOL*) - -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = HOL.False -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = HOL.True; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun in_interval (k, l) n = - HOL.anda (Nat.less_eq_nat k n) (Nat.less_eq_nat n l); - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/bool_mlbool.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/bool_mlbool.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun in_interval (k, l) n = - (Nat.less_eq_nat k n) andalso (Nat.less_eq_nat n l); - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/class.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/class.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -end; (*struct Nat*) - -structure Codegen = -struct - -type 'a null = {null : 'a}; -fun null (A_:'a null) = #null A_; - -fun head A_ (x :: xs) = x - | head A_ [] = null A_; - -val null_option : 'a option = NONE; - -fun null_optiona () = {null = null_option} : ('a option) null; - -val dummy : Nat.nat option = - head (null_optiona ()) [SOME (Nat.Suc Nat.Zero_nat), NONE]; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/class.ocaml --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/class.ocaml Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ -module Nat = -struct - -type nat = Suc of nat | Zero_nat;; - -end;; (*struct Nat*) - -module Codegen = -struct - -type 'a null = {null : 'a};; -let null _A = _A.null;; - -let rec head _A = function x :: xs -> x - | [] -> null _A;; - -let rec null_option = None;; - -let null_optiona () = ({null = null_option} : ('a option) null);; - -let rec dummy - = head (null_optiona ()) [Some (Nat.Suc Nat.Zero_nat); None];; - -end;; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/collect_duplicates.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/collect_duplicates.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,30 +0,0 @@ -structure HOL = -struct - -type 'a eq = {eq : 'a -> 'a -> bool}; -fun eq (A_:'a eq) = #eq A_; - -fun eqop A_ a = eq A_ a; - -end; (*struct HOL*) - -structure List = -struct - -fun member A_ x (y :: ys) = - (if HOL.eqop A_ y x then true else member A_ x ys) - | member A_ x [] = false; - -end; (*struct List*) - -structure Codegen = -struct - -fun collect_duplicates A_ xs ys (z :: zs) = - (if List.member A_ z xs - then (if List.member A_ z ys then collect_duplicates A_ xs ys zs - else collect_duplicates A_ xs (z :: ys) zs) - else collect_duplicates A_ (z :: xs) (z :: ys) zs) - | collect_duplicates A_ xs ys [] = xs; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/dirty_set.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/dirty_set.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,102 +0,0 @@ -structure ROOT = -struct - -structure Nat = -struct - -datatype nat = Zero_nat | Suc of nat; - -end; (*struct Nat*) - -structure Integer = -struct - -datatype bit = B0 | B1; - -datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; - -fun pred (Bit (k, B0)) = Bit (pred k, B1) - | pred (Bit (k, B1)) = Bit (k, B0) - | pred Min = Bit (Min, B0) - | pred Pls = Min; - -fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) - | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) - | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) - | uminus_int Min = Bit (Pls, B1) - | uminus_int Pls = Pls; - -fun succ (Bit (k, B0)) = Bit (k, B1) - | succ (Bit (k, B1)) = Bit (succ k, B0) - | succ Min = Pls - | succ Pls = Bit (Pls, B1); - -fun plus_int (Number_of_int v) (Number_of_int w) = - Number_of_int (plus_int v w) - | plus_int k Min = pred k - | plus_int k Pls = k - | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) - | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) - | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) - | plus_int Min k = pred k - | plus_int Pls k = k; - -fun minus_int (Number_of_int v) (Number_of_int w) = - Number_of_int (plus_int v (uminus_int w)) - | minus_int z w = plus_int z (uminus_int w); - -fun less_eq_int (Number_of_int k) (Number_of_int l) = less_eq_int k l - | less_eq_int (Bit (k1, B1)) (Bit (k2, B0)) = less_int k1 k2 - | less_eq_int (Bit (k1, v)) (Bit (k2, B1)) = less_eq_int k1 k2 - | less_eq_int (Bit (k1, B0)) (Bit (k2, v)) = less_eq_int k1 k2 - | less_eq_int (Bit (k, v)) Min = less_eq_int k Min - | less_eq_int (Bit (k, B1)) Pls = less_int k Pls - | less_eq_int (Bit (k, B0)) Pls = less_eq_int k Pls - | less_eq_int Min (Bit (k, B1)) = less_eq_int Min k - | less_eq_int Min (Bit (k, B0)) = less_int Min k - | less_eq_int Min Min = true - | less_eq_int Min Pls = true - | less_eq_int Pls (Bit (k, v)) = less_eq_int Pls k - | less_eq_int Pls Min = false - | less_eq_int Pls Pls = true -and less_int (Number_of_int k) (Number_of_int l) = less_int k l - | less_int (Bit (k1, B0)) (Bit (k2, B1)) = less_eq_int k1 k2 - | less_int (Bit (k1, B1)) (Bit (k2, v)) = less_int k1 k2 - | less_int (Bit (k1, v)) (Bit (k2, B0)) = less_int k1 k2 - | less_int (Bit (k, B1)) Min = less_int k Min - | less_int (Bit (k, B0)) Min = less_eq_int k Min - | less_int (Bit (k, v)) Pls = less_int k Pls - | less_int Min (Bit (k, v)) = less_int Min k - | less_int Min Min = false - | less_int Min Pls = true - | less_int Pls (Bit (k, B1)) = less_eq_int Pls k - | less_int Pls (Bit (k, B0)) = less_int Pls k - | less_int Pls Min = false - | less_int Pls Pls = false; - -fun nat_aux n i = - (if less_eq_int i (Number_of_int Pls) then n - else nat_aux (Nat.Suc n) - (minus_int i (Number_of_int (Bit (Pls, B1))))); - -fun nat i = nat_aux Nat.Zero_nat i; - -end; (*struct Integer*) - -structure Codegen = -struct - -val dummy_set : (Nat.nat -> Nat.nat) list = Nat.Suc :: []; - -val foobar_set : Nat.nat list = - Nat.Zero_nat :: - (Nat.Suc Nat.Zero_nat :: - (Integer.nat - (Integer.Number_of_int - (Integer.Bit - (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) - :: [])); - -end; (*struct Codegen*) - -end; (*struct ROOT*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,27 +0,0 @@ -structure Example = -struct - -fun foldl f a [] = a - | foldl f a (x :: xs) = foldl f (f a x) xs; - -fun rev xs = foldl (fn xsa => fn x => x :: xsa) [] xs; - -fun list_case f1 f2 (a :: lista) = f2 a lista - | list_case f1 f2 [] = f1; - -datatype 'a queue = AQueue of 'a list * 'a list; - -val empty : 'a queue = AQueue ([], []) - -fun dequeue (AQueue ([], [])) = (NONE, AQueue ([], [])) - | dequeue (AQueue (xs, y :: ys)) = (SOME y, AQueue (xs, ys)) - | dequeue (AQueue (v :: va, [])) = - let - val y :: ys = rev (v :: va); - in - (SOME y, AQueue ([], ys)) - end; - -fun enqueue x (AQueue (xs, ys)) = AQueue (x :: xs, ys); - -end; (*struct Example*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/fac.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/fac.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,22 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -val one_nat : nat = Suc Zero_nat; - -fun plus_nat (Suc m) n = plus_nat m (Suc n) - | plus_nat Zero_nat n = n; - -fun times_nat (Suc m) n = plus_nat n (times_nat m n) - | times_nat Zero_nat n = Zero_nat; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun fac (Nat.Suc n) = Nat.times_nat (Nat.Suc n) (fac n) - | fac Nat.Zero_nat = Nat.one_nat; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/integers.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/integers.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,59 +0,0 @@ -structure ROOT = -struct - -structure Integer = -struct - -datatype bit = B0 | B1; - -datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; - -fun pred (Bit (k, B0)) = Bit (pred k, B1) - | pred (Bit (k, B1)) = Bit (k, B0) - | pred Min = Bit (Min, B0) - | pred Pls = Min; - -fun succ (Bit (k, B0)) = Bit (k, B1) - | succ (Bit (k, B1)) = Bit (succ k, B0) - | succ Min = Pls - | succ Pls = Bit (Pls, B1); - -fun plus_int (Number_of_int v) (Number_of_int w) = - Number_of_int (plus_int v w) - | plus_int k Min = pred k - | plus_int k Pls = k - | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) - | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) - | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) - | plus_int Min k = pred k - | plus_int Pls k = k; - -fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) - | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) - | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) - | uminus_int Min = Bit (Pls, B1) - | uminus_int Pls = Pls; - -fun times_int (Number_of_int v) (Number_of_int w) = - Number_of_int (times_int v w) - | times_int (Bit (k, B0)) l = Bit (times_int k l, B0) - | times_int (Bit (k, B1)) l = plus_int (Bit (times_int k l, B0)) l - | times_int Min k = uminus_int k - | times_int Pls w = Pls; - -end; (*struct Integer*) - -structure Codegen = -struct - -fun double_inc k = - Integer.plus_int - (Integer.times_int - (Integer.Number_of_int - (Integer.Bit (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) - k) - (Integer.Number_of_int (Integer.Bit (Integer.Pls, Integer.B1))); - -end; (*struct Codegen*) - -end; (*struct ROOT*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/lexicographic.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/lexicographic.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -structure HOL = -struct - -type 'a eq = {eq : 'a -> 'a -> bool}; -fun eq (A_:'a eq) = #eq A_; - -type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; -fun less_eq (A_:'a ord) = #less_eq A_; -fun less (A_:'a ord) = #less A_; - -end; (*struct HOL*) - -structure Codegen = -struct - -fun less_eq (A1_, A2_) B_ (x1, y1) (x2, y2) = - HOL.less A2_ x1 x2 orelse HOL.eq A1_ x1 x2 andalso HOL.less_eq B_ y1 y2; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/lookup.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/lookup.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,13 +0,0 @@ -structure ROOT = -struct - -structure Codegen = -struct - -fun lookup ((k, v) :: xs) l = - (if ((k : string) = l) then SOME v else lookup xs l) - | lookup [] l = NONE; - -end; (*struct Codegen*) - -end; (*struct ROOT*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/monotype.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/monotype.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,34 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun eq_nat (Suc a) Zero_nat = false - | eq_nat Zero_nat (Suc a) = false - | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' - | eq_nat Zero_nat Zero_nat = true; - -end; (*struct Nat*) - -structure List = -struct - -fun null (x :: xs) = false - | null [] = true; - -fun list_all2 p (x :: xs) (y :: ys) = p x y andalso list_all2 p xs ys - | list_all2 p xs [] = null xs - | list_all2 p [] ys = null ys; - -end; (*struct List*) - -structure Codegen = -struct - -datatype monotype = Mono of Nat.nat * monotype list; - -fun eq_monotype (Mono (tyco1, typargs1)) (Mono (tyco2, typargs2)) = - Nat.eq_nat tyco1 tyco2 andalso - List.list_all2 eq_monotype typargs1 typargs2; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/nat_binary.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/nat_binary.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,17 +0,0 @@ -structure Nat = -struct - -datatype nat = Dig1 of nat | Dig0 of nat | One_nat | Zero_nat; - -fun plus_nat (Dig1 m) (Dig1 n) = Dig0 (plus_nat (plus_nat m n) One_nat) - | plus_nat (Dig1 m) (Dig0 n) = Dig1 (plus_nat m n) - | plus_nat (Dig0 m) (Dig1 n) = Dig1 (plus_nat m n) - | plus_nat (Dig0 m) (Dig0 n) = Dig0 (plus_nat m n) - | plus_nat (Dig1 m) One_nat = Dig0 (plus_nat m One_nat) - | plus_nat One_nat (Dig1 n) = Dig0 (plus_nat n One_nat) - | plus_nat (Dig0 m) One_nat = Dig1 m - | plus_nat One_nat (Dig0 n) = Dig1 n - | plus_nat m Zero_nat = m - | plus_nat Zero_nat n = n; - -end; (*struct Nat*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/pick1.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/pick1.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,44 +0,0 @@ -structure HOL = -struct - -fun leta s f = f s; - -end; (*struct HOL*) - -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -fun minus_nat (Suc m) (Suc n) = minus_nat m n - | minus_nat Zero_nat n = Zero_nat - | minus_nat m Zero_nat = m; - -end; (*struct Nat*) - -structure Product_Type = -struct - -fun split f (a, b) = f a b; - -end; (*struct Product_Type*) - -structure Codegen = -struct - -fun pick ((k, v) :: xs) n = - (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) - | pick (x :: xs) n = - let - val a = x; - val (k, v) = a; - in - (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) - end; - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/Thy/examples/tree.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/tree.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,95 +0,0 @@ -structure HOL = -struct - -type 'a eq = {eq : 'a -> 'a -> bool}; -fun eq (A_:'a eq) = #eq A_; - -type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; -fun less_eq (A_:'a ord) = #less_eq A_; -fun less (A_:'a ord) = #less A_; - -fun eqop A_ a = eq A_ a; - -end; (*struct HOL*) - -structure Orderings = -struct - -type 'a preorder = {Orderings__ord_preorder : 'a HOL.ord}; -fun ord_preorder (A_:'a preorder) = #Orderings__ord_preorder A_; - -type 'a order = {Orderings__preorder_order : 'a preorder}; -fun preorder_order (A_:'a order) = #Orderings__preorder_order A_; - -type 'a linorder = {Orderings__order_linorder : 'a order}; -fun order_linorder (A_:'a linorder) = #Orderings__order_linorder A_; - -end; (*struct Orderings*) - -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun eq_nat (Suc a) Zero_nat = false - | eq_nat Zero_nat (Suc a) = false - | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' - | eq_nat Zero_nat Zero_nat = true; - -val eq_nata = {eq = eq_nat} : nat HOL.eq; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -val ord_nat = {less_eq = less_eq_nat, less = less_nat} : nat HOL.ord; - -val preorder_nat = {Orderings__ord_preorder = ord_nat} : - nat Orderings.preorder; - -val order_nat = {Orderings__preorder_order = preorder_nat} : - nat Orderings.order; - -val linorder_nat = {Orderings__order_linorder = order_nat} : - nat Orderings.linorder; - -end; (*struct Nat*) - -structure Codegen = -struct - -datatype ('a, 'b) searchtree = - Branch of ('a, 'b) searchtree * 'a * ('a, 'b) searchtree | - Leaf of 'a * 'b; - -fun update (A1_, A2_) (it, entry) (Leaf (key, vala)) = - (if HOL.eqop A1_ it key then Leaf (key, entry) - else (if HOL.less_eq - ((Orderings.ord_preorder o Orderings.preorder_order o - Orderings.order_linorder) - A2_) - it key - then Branch (Leaf (it, entry), it, Leaf (key, vala)) - else Branch (Leaf (key, vala), it, Leaf (it, entry)))) - | update (A1_, A2_) (it, entry) (Branch (t1, key, t2)) = - (if HOL.less_eq - ((Orderings.ord_preorder o Orderings.preorder_order o - Orderings.order_linorder) - A2_) - it key - then Branch (update (A1_, A2_) (it, entry) t1, key, t2) - else Branch (t1, key, update (A1_, A2_) (it, entry) t2)); - -val example : (Nat.nat, (Nat.nat list)) searchtree = - update (Nat.eq_nata, Nat.linorder_nat) - (Nat.Suc (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))), - [Nat.Suc (Nat.Suc Nat.Zero_nat), Nat.Suc (Nat.Suc Nat.Zero_nat)]) - (update (Nat.eq_nata, Nat.linorder_nat) - (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat)), - [Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))]) - (update (Nat.eq_nata, Nat.linorder_nat) - (Nat.Suc (Nat.Suc Nat.Zero_nat), [Nat.Suc (Nat.Suc Nat.Zero_nat)]) - (Leaf (Nat.Suc Nat.Zero_nat, [])))); - -end; (*struct Codegen*) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/codegen.tex --- a/doc-src/IsarAdvanced/Codegen/codegen.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,53 +0,0 @@ - -\documentclass[12pt,a4paper,fleqn]{report} -\usepackage{latexsym,graphicx} -\usepackage[refpage]{nomencl} -\usepackage{../../iman,../../extra,../../isar,../../proof} -\usepackage{../../isabelle,../../isabellesym} -\usepackage{style} -\usepackage{pgf} -\usepackage{pgflibraryshapes} -\usepackage{tikz} -\usepackage{../../pdfsetup} - -\hyphenation{Isabelle} -\hyphenation{Isar} -\isadroptag{theory} - -\title{\includegraphics[scale=0.5]{isabelle_isar} - \\[4ex] Code generation from Isabelle/HOL theories} -\author{\emph{Florian Haftmann}} - -\begin{document} - -\maketitle - -\begin{abstract} - This tutorial gives an introduction to a generic code generator framework in Isabelle - for generating executable code in functional programming languages from logical - specifications in Isabelle/HOL. -\end{abstract} - -\thispagestyle{empty}\clearpage - -\pagenumbering{roman} -\clearfirst - -\input{Thy/document/Introduction.tex} -\input{Thy/document/Program.tex} -\input{Thy/document/Adaption.tex} -\input{Thy/document/Further.tex} -\input{Thy/document/ML.tex} - -\begingroup -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{../../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/codegen_process.pdf Binary file doc-src/IsarAdvanced/Codegen/codegen_process.pdf has changed diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/codegen_process.ps --- a/doc-src/IsarAdvanced/Codegen/codegen_process.ps Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,586 +0,0 @@ -%!PS-Adobe-2.0 -%%Creator: dot version 2.2 (Mon Sep 12 23:33:36 UTC 2005) -%%For: (haftmann) Florian Haftmann -%%Title: _anonymous_0 -%%Pages: (atend) -%%BoundingBox: 35 35 451 291 -%%EndComments -save -%%BeginProlog -/DotDict 200 dict def -DotDict begin - -/setupLatin1 { -mark -/EncodingVector 256 array def - EncodingVector 0 - -ISOLatin1Encoding 0 255 getinterval putinterval - -EncodingVector - dup 306 /AE - dup 301 /Aacute - dup 302 /Acircumflex - dup 304 /Adieresis - dup 300 /Agrave - dup 305 /Aring - dup 303 /Atilde - dup 307 /Ccedilla - dup 311 /Eacute - dup 312 /Ecircumflex - dup 313 /Edieresis - dup 310 /Egrave - dup 315 /Iacute - dup 316 /Icircumflex - dup 317 /Idieresis - dup 314 /Igrave - dup 334 /Udieresis - dup 335 /Yacute - dup 376 /thorn - dup 337 /germandbls - dup 341 /aacute - dup 342 /acircumflex - dup 344 /adieresis - dup 346 /ae - dup 340 /agrave - dup 345 /aring - dup 347 /ccedilla - dup 351 /eacute - dup 352 /ecircumflex - dup 353 /edieresis - dup 350 /egrave - dup 355 /iacute - dup 356 /icircumflex - dup 357 /idieresis - dup 354 /igrave - dup 360 /dcroat - dup 361 /ntilde - dup 363 /oacute - dup 364 /ocircumflex - dup 366 /odieresis - dup 362 /ograve - dup 365 /otilde - dup 370 /oslash - dup 372 /uacute - dup 373 /ucircumflex - dup 374 /udieresis - dup 371 /ugrave - dup 375 /yacute - dup 377 /ydieresis - -% Set up ISO Latin 1 character encoding -/starnetISO { - dup dup findfont dup length dict begin - { 1 index /FID ne { def }{ pop pop } ifelse - } forall - /Encoding EncodingVector def - currentdict end definefont -} def -/Times-Roman starnetISO def -/Times-Italic starnetISO def -/Times-Bold starnetISO def -/Times-BoldItalic starnetISO def -/Helvetica starnetISO def -/Helvetica-Oblique starnetISO def -/Helvetica-Bold starnetISO def -/Helvetica-BoldOblique starnetISO def -/Courier starnetISO def -/Courier-Oblique starnetISO def -/Courier-Bold starnetISO def -/Courier-BoldOblique starnetISO def -cleartomark -} bind def - -%%BeginResource: procset graphviz 0 0 -/coord-font-family /Times-Roman def -/default-font-family /Times-Roman def -/coordfont coord-font-family findfont 8 scalefont def - -/InvScaleFactor 1.0 def -/set_scale { - dup 1 exch div /InvScaleFactor exch def - dup scale -} bind def - -% styles -/solid { [] 0 setdash } bind def -/dashed { [9 InvScaleFactor mul dup ] 0 setdash } bind def -/dotted { [1 InvScaleFactor mul 6 InvScaleFactor mul] 0 setdash } bind def -/invis {/fill {newpath} def /stroke {newpath} def /show {pop newpath} def} bind def -/bold { 2 setlinewidth } bind def -/filled { } bind def -/unfilled { } bind def -/rounded { } bind def -/diagonals { } bind def - -% hooks for setting color -/nodecolor { sethsbcolor } bind def -/edgecolor { sethsbcolor } bind def -/graphcolor { sethsbcolor } bind def -/nopcolor {pop pop pop} bind def - -/beginpage { % i j npages - /npages exch def - /j exch def - /i exch def - /str 10 string def - npages 1 gt { - gsave - coordfont setfont - 0 0 moveto - (\() show i str cvs show (,) show j str cvs show (\)) show - grestore - } if -} bind def - -/set_font { - findfont exch - scalefont setfont -} def - -% draw aligned label in bounding box aligned to current point -/alignedtext { % width adj text - /text exch def - /adj exch def - /width exch def - gsave - width 0 gt { - text stringwidth pop adj mul 0 rmoveto - } if - [] 0 setdash - text show - grestore -} def - -/boxprim { % xcorner ycorner xsize ysize - 4 2 roll - moveto - 2 copy - exch 0 rlineto - 0 exch rlineto - pop neg 0 rlineto - closepath -} bind def - -/ellipse_path { - /ry exch def - /rx exch def - /y exch def - /x exch def - matrix currentmatrix - newpath - x y translate - rx ry scale - 0 0 1 0 360 arc - setmatrix -} bind def - -/endpage { showpage } bind def -/showpage { } def - -/layercolorseq - [ % layer color sequence - darkest to lightest - [0 0 0] - [.2 .8 .8] - [.4 .8 .8] - [.6 .8 .8] - [.8 .8 .8] - ] -def - -/layerlen layercolorseq length def - -/setlayer {/maxlayer exch def /curlayer exch def - layercolorseq curlayer 1 sub layerlen mod get - aload pop sethsbcolor - /nodecolor {nopcolor} def - /edgecolor {nopcolor} def - /graphcolor {nopcolor} def -} bind def - -/onlayer { curlayer ne {invis} if } def - -/onlayers { - /myupper exch def - /mylower exch def - curlayer mylower lt - curlayer myupper gt - or - {invis} if -} def - -/curlayer 0 def - -%%EndResource -%%EndProlog -%%BeginSetup -14 default-font-family set_font -1 setmiterlimit -% /arrowlength 10 def -% /arrowwidth 5 def - -% make sure pdfmark is harmless for PS-interpreters other than Distiller -/pdfmark where {pop} {userdict /pdfmark /cleartomark load put} ifelse -% make '<<' and '>>' safe on PS Level 1 devices -/languagelevel where {pop languagelevel}{1} ifelse -2 lt { - userdict (<<) cvn ([) cvn load put - userdict (>>) cvn ([) cvn load put -} if - -%%EndSetup -%%Page: 1 1 -%%PageBoundingBox: 36 36 451 291 -%%PageOrientation: Portrait -gsave -35 35 416 256 boxprim clip newpath -36 36 translate -0 0 1 beginpage -0 0 translate 0 rotate -[ /CropBox [36 36 451 291] /PAGES pdfmark -0.000 0.000 0.000 graphcolor -14.00 /Times-Roman set_font - -% theory -gsave 10 dict begin -newpath 93 254 moveto -1 254 lineto -1 214 lineto -93 214 lineto -closepath -stroke -gsave 10 dict begin -8 237 moveto -(Isabelle/HOL) -[4.56 5.52 6.24 6.96 6.24 3.84 3.84 6.24 3.84 10.08 10.08 8.64] -xshow -16 221 moveto -(Isar theory) -[4.56 5.52 6.24 4.56 3.6 4.08 6.96 6.24 6.96 4.8 6.96] -xshow -end grestore -end grestore - -% selection -gsave 10 dict begin -183 234 38 18 ellipse_path -stroke -gsave 10 dict begin -158 229 moveto -(selection) -[5.52 6.24 3.84 6.24 6.24 3.84 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% theory -> selection -newpath 94 234 moveto -107 234 121 234 135 234 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 135 238 moveto -145 234 lineto -135 231 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 135 238 moveto -145 234 lineto -135 231 lineto -closepath -stroke -end grestore - -% sml -gsave 10 dict begin -newpath 74 144 moveto -20 144 lineto -20 108 lineto -74 108 lineto -closepath -stroke -gsave 10 dict begin -32 121 moveto -(SML) -[7.68 12.48 8.64] -xshow -end grestore -end grestore - -% other -gsave 10 dict begin -gsave 10 dict begin -41 67 moveto -(...) -[3.6 3.6 3.6] -xshow -end grestore -end grestore - -% haskell -gsave 10 dict begin -newpath 77 36 moveto -17 36 lineto -17 0 lineto -77 0 lineto -closepath -stroke -gsave 10 dict begin -25 13 moveto -(Haskell) -[10.08 6.24 5.52 6.72 6.24 3.84 3.84] -xshow -end grestore -end grestore - -% preprocessing -gsave 10 dict begin -183 180 52 18 ellipse_path -stroke -gsave 10 dict begin -143 175 moveto -(preprocessing) -[6.96 4.56 6.24 6.96 4.56 6.96 6.24 6.24 5.52 5.52 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% selection -> preprocessing -newpath 183 216 moveto -183 213 183 211 183 208 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 187 208 moveto -183 198 lineto -180 208 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 187 208 moveto -183 198 lineto -180 208 lineto -closepath -stroke -end grestore - -% def_eqn -gsave 10 dict begin -newpath 403 198 moveto -283 198 lineto -283 162 lineto -403 162 lineto -closepath -stroke -gsave 10 dict begin -291 175 moveto -(defining equations) -[6.96 6.24 4.8 3.84 6.96 3.84 6.96 6.96 3.6 6.24 6.72 6.96 6.24 3.84 3.84 6.96 6.96 5.52] -xshow -end grestore -end grestore - -% preprocessing -> def_eqn -newpath 236 180 moveto -248 180 260 180 273 180 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 273 184 moveto -283 180 lineto -273 177 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 273 184 moveto -283 180 lineto -273 177 lineto -closepath -stroke -end grestore - -% serialization -gsave 10 dict begin -183 72 47 18 ellipse_path -stroke -gsave 10 dict begin -148 67 moveto -(serialization) -[5.52 6.24 4.8 3.84 6.24 3.84 3.84 6.24 6.24 3.84 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% serialization -> sml -newpath 150 85 moveto -129 93 104 103 83 111 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 82 108 moveto -74 115 lineto -85 114 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 82 108 moveto -74 115 lineto -85 114 lineto -closepath -stroke -end grestore - -% serialization -> other -gsave 10 dict begin -dotted -newpath 135 72 moveto -119 72 100 72 84 72 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 84 69 moveto -74 72 lineto -84 76 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 84 69 moveto -74 72 lineto -84 76 lineto -closepath -stroke -end grestore -end grestore - -% serialization -> haskell -newpath 150 59 moveto -131 51 107 42 86 34 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 88 31 moveto -77 30 lineto -85 37 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 88 31 moveto -77 30 lineto -85 37 lineto -closepath -stroke -end grestore - -% translation -gsave 10 dict begin -343 126 43 18 ellipse_path -stroke -gsave 10 dict begin -313 121 moveto -(translation) -[3.84 4.56 6.24 6.96 5.52 3.84 6.24 3.84 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% def_eqn -> translation -newpath 343 162 moveto -343 159 343 157 343 154 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 347 154 moveto -343 144 lineto -340 154 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 347 154 moveto -343 144 lineto -340 154 lineto -closepath -stroke -end grestore - -% iml -gsave 10 dict begin -newpath 413 90 moveto -273 90 lineto -273 54 lineto -413 54 lineto -closepath -stroke -gsave 10 dict begin -280 67 moveto -(intermediate language) -[3.84 6.96 3.84 6.24 4.8 10.8 6.24 6.96 3.84 6.24 3.84 6.24 3.6 3.84 6.24 6.96 6.96 6.96 6.24 6.72 6.24] -xshow -end grestore -end grestore - -% translation -> iml -newpath 343 108 moveto -343 105 343 103 343 100 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 347 100 moveto -343 90 lineto -340 100 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 347 100 moveto -343 90 lineto -340 100 lineto -closepath -stroke -end grestore - -% iml -> serialization -newpath 272 72 moveto -262 72 251 72 241 72 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 241 69 moveto -231 72 lineto -241 76 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 241 69 moveto -231 72 lineto -241 76 lineto -closepath -stroke -end grestore -endpage -showpage -grestore -%%PageTrailer -%%EndPage: 1 -%%Trailer -%%Pages: 1 -end -restore -%%EOF diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Codegen/style.sty --- a/doc-src/IsarAdvanced/Codegen/style.sty Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,56 +0,0 @@ - -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} - -%% logical markup -\newcommand{\strong}[1]{{\bfseries {#1}}} -\newcommand{\qn}[1]{\emph{#1}} - -%% typographic conventions -\newcommand{\qt}[1]{``{#1}''} - -%% verbatim text -\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} - -%% quoted segments -\makeatletter -\isakeeptag{quote} -\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}} -\renewcommand{\isatagquote}{\begin{quotesegment}} -\renewcommand{\endisatagquote}{\end{quotesegment}} -\makeatother - -\isakeeptag{quotett} -\renewcommand{\isatagquotett}{\begin{quotesegment}\isabellestyle{tt}\isastyle} -\renewcommand{\endisatagquotett}{\end{quotesegment}\isabellestyle{it}\isastyle} - -%% a trick -\newcommand{\isasymSML}{SML} - -%% presentation -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\binperiod -\underscoreoff - -\renewcommand{\isadigit}[1]{\isamath{#1}} - -%% ml reference -\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} - -\isakeeptag{mlref} -\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} -\renewcommand{\endisatagmlref}{\endgroup} - -\isabellestyle{it} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/IsaMakefile --- a/doc-src/IsarAdvanced/Functions/IsaMakefile Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ - -## targets - -default: Thy -images: -test: Thy - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log - -USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document - - -## Thy - -THY = $(LOG)/HOL-Thy.gz - -Thy: $(THY) - -$(THY): Thy/ROOT.ML Thy/Functions.thy - @$(USEDIR) HOL Thy - - -## clean - -clean: - @rm -f $(THY) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/Makefile --- a/doc-src/IsarAdvanced/Functions/Makefile Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = functions - -FILES = $(NAME).tex Thy/document/Functions.tex intro.tex conclusion.tex \ - style.sty ../../iman.sty ../../extra.sty ../../isar.sty \ - ../../isabelle.sty ../../isabellesym.sty ../../pdfsetup.sty \ - ../../manual.bib ../../proof.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf - $(PDFLATEX) $(NAME) - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/Thy/Functions.thy --- a/doc-src/IsarAdvanced/Functions/Thy/Functions.thy Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1264 +0,0 @@ -(* Title: doc-src/IsarAdvanced/Functions/Thy/Fundefs.thy - Author: Alexander Krauss, TU Muenchen - -Tutorial for function definitions with the new "function" package. -*) - -theory Functions -imports Main -begin - -section {* Function Definitions for Dummies *} - -text {* - In most cases, defining a recursive function is just as simple as other definitions: -*} - -fun fib :: "nat \ nat" -where - "fib 0 = 1" -| "fib (Suc 0) = 1" -| "fib (Suc (Suc n)) = fib n + fib (Suc n)" - -text {* - The syntax is rather self-explanatory: We introduce a function by - giving its name, its type, - and a set of defining recursive equations. - If we leave out the type, the most general type will be - inferred, which can sometimes lead to surprises: Since both @{term - "1::nat"} and @{text "+"} are overloaded, we would end up - with @{text "fib :: nat \ 'a::{one,plus}"}. -*} - -text {* - The function always terminates, since its argument gets smaller in - every recursive call. - Since HOL is a logic of total functions, termination is a - fundamental requirement to prevent inconsistencies\footnote{From the - \qt{definition} @{text "f(n) = f(n) + 1"} we could prove - @{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.}. - Isabelle tries to prove termination automatically when a definition - is made. In \S\ref{termination}, we will look at cases where this - fails and see what to do then. -*} - -subsection {* Pattern matching *} - -text {* \label{patmatch} - Like in functional programming, we can use pattern matching to - define functions. At the moment we will only consider \emph{constructor - patterns}, which only consist of datatype constructors and - variables. Furthermore, patterns must be linear, i.e.\ all variables - on the left hand side of an equation must be distinct. In - \S\ref{genpats} we discuss more general pattern matching. - - If patterns overlap, the order of the equations is taken into - account. The following function inserts a fixed element between any - two elements of a list: -*} - -fun sep :: "'a \ 'a list \ 'a list" -where - "sep a (x#y#xs) = x # a # sep a (y # xs)" -| "sep a xs = xs" - -text {* - Overlapping patterns are interpreted as \qt{increments} to what is - already there: The second equation is only meant for the cases where - the first one does not match. Consequently, Isabelle replaces it - internally by the remaining cases, making the patterns disjoint: -*} - -thm sep.simps - -text {* @{thm [display] sep.simps[no_vars]} *} - -text {* - \noindent The equations from function definitions are automatically used in - simplification: -*} - -lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]" -by simp - -subsection {* Induction *} - -text {* - - Isabelle provides customized induction rules for recursive - functions. These rules follow the recursive structure of the - definition. Here is the rule @{text sep.induct} arising from the - above definition of @{const sep}: - - @{thm [display] sep.induct} - - We have a step case for list with at least two elements, and two - base cases for the zero- and the one-element list. Here is a simple - proof about @{const sep} and @{const map} -*} - -lemma "map f (sep x ys) = sep (f x) (map f ys)" -apply (induct x ys rule: sep.induct) - -txt {* - We get three cases, like in the definition. - - @{subgoals [display]} -*} - -apply auto -done -text {* - - With the \cmd{fun} command, you can define about 80\% of the - functions that occur in practice. The rest of this tutorial explains - the remaining 20\%. -*} - - -section {* fun vs.\ function *} - -text {* - The \cmd{fun} command provides a - convenient shorthand notation for simple function definitions. In - this mode, Isabelle tries to solve all the necessary proof obligations - automatically. If any proof fails, the definition is - rejected. This can either mean that the definition is indeed faulty, - or that the default proof procedures are just not smart enough (or - rather: not designed) to handle the definition. - - By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or - solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: - -\end{isamarkuptext} - - -\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} -\cmd{fun} @{text "f :: \"}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\vspace*{6pt} -\end{minipage}\right] -\quad\equiv\quad -\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} -\cmd{function} @{text "("}\cmd{sequential}@{text ") f :: \"}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\\% -\cmd{by} @{text "pat_completeness auto"}\\% -\cmd{termination by} @{text "lexicographic_order"}\vspace{6pt} -\end{minipage} -\right]\] - -\begin{isamarkuptext} - \vspace*{1em} - \noindent Some details have now become explicit: - - \begin{enumerate} - \item The \cmd{sequential} option enables the preprocessing of - pattern overlaps which we already saw. Without this option, the equations - must already be disjoint and complete. The automatic completion only - works with constructor patterns. - - \item A function definition produces a proof obligation which - expresses completeness and compatibility of patterns (we talk about - this later). The combination of the methods @{text "pat_completeness"} and - @{text "auto"} is used to solve this proof obligation. - - \item A termination proof follows the definition, started by the - \cmd{termination} command. This will be explained in \S\ref{termination}. - \end{enumerate} - Whenever a \cmd{fun} command fails, it is usually a good idea to - expand the syntax to the more verbose \cmd{function} form, to see - what is actually going on. - *} - - -section {* Termination *} - -text {*\label{termination} - The method @{text "lexicographic_order"} is the default method for - termination proofs. It can prove termination of a - certain class of functions by searching for a suitable lexicographic - combination of size measures. Of course, not all functions have such - a simple termination argument. For them, we can specify the termination - relation manually. -*} - -subsection {* The {\tt relation} method *} -text{* - Consider the following function, which sums up natural numbers up to - @{text "N"}, using a counter @{text "i"}: -*} - -function sum :: "nat \ nat \ nat" -where - "sum i N = (if i > N then 0 else i + sum (Suc i) N)" -by pat_completeness auto - -text {* - \noindent The @{text "lexicographic_order"} method fails on this example, because none of the - arguments decreases in the recursive call, with respect to the standard size ordering. - To prove termination manually, we must provide a custom wellfounded relation. - - The termination argument for @{text "sum"} is based on the fact that - the \emph{difference} between @{text "i"} and @{text "N"} gets - smaller in every step, and that the recursion stops when @{text "i"} - is greater than @{text "N"}. Phrased differently, the expression - @{text "N + 1 - i"} always decreases. - - We can use this expression as a measure function suitable to prove termination. -*} - -termination sum -apply (relation "measure (\(i,N). N + 1 - i)") - -txt {* - The \cmd{termination} command sets up the termination goal for the - specified function @{text "sum"}. If the function name is omitted, it - implicitly refers to the last function definition. - - The @{text relation} method takes a relation of - type @{typ "('a \ 'a) set"}, where @{typ "'a"} is the argument type of - the function. If the function has multiple curried arguments, then - these are packed together into a tuple, as it happened in the above - example. - - The predefined function @{term[source] "measure :: ('a \ nat) \ ('a \ 'a) set"} constructs a - wellfounded relation from a mapping into the natural numbers (a - \emph{measure function}). - - After the invocation of @{text "relation"}, we must prove that (a) - the relation we supplied is wellfounded, and (b) that the arguments - of recursive calls indeed decrease with respect to the - relation: - - @{subgoals[display,indent=0]} - - These goals are all solved by @{text "auto"}: -*} - -apply auto -done - -text {* - Let us complicate the function a little, by adding some more - recursive calls: -*} - -function foo :: "nat \ nat \ nat" -where - "foo i N = (if i > N - then (if N = 0 then 0 else foo 0 (N - 1)) - else i + foo (Suc i) N)" -by pat_completeness auto - -text {* - When @{text "i"} has reached @{text "N"}, it starts at zero again - and @{text "N"} is decremented. - This corresponds to a nested - loop where one index counts up and the other down. Termination can - be proved using a lexicographic combination of two measures, namely - the value of @{text "N"} and the above difference. The @{const - "measures"} combinator generalizes @{text "measure"} by taking a - list of measure functions. -*} - -termination -by (relation "measures [\(i, N). N, \(i,N). N + 1 - i]") auto - -subsection {* How @{text "lexicographic_order"} works *} - -(*fun fails :: "nat \ nat list \ nat" -where - "fails a [] = a" -| "fails a (x#xs) = fails (x + a) (x # xs)" -*) - -text {* - To see how the automatic termination proofs work, let's look at an - example where it fails\footnote{For a detailed discussion of the - termination prover, see \cite{bulwahnKN07}}: - -\end{isamarkuptext} -\cmd{fun} @{text "fails :: \"nat \ nat list \ nat\""}\\% -\cmd{where}\\% -\hspace*{2ex}@{text "\"fails a [] = a\""}\\% -|\hspace*{1.5ex}@{text "\"fails a (x#xs) = fails (x + a) (x#xs)\""}\\ -\begin{isamarkuptext} - -\noindent Isabelle responds with the following error: - -\begin{isabelle} -*** Unfinished subgoals:\newline -*** (a, 1, <):\newline -*** \ 1.~@{text "\x. x = 0"}\newline -*** (a, 1, <=):\newline -*** \ 1.~False\newline -*** (a, 2, <):\newline -*** \ 1.~False\newline -*** Calls:\newline -*** a) @{text "(a, x # xs) -->> (x + a, x # xs)"}\newline -*** Measures:\newline -*** 1) @{text "\x. size (fst x)"}\newline -*** 2) @{text "\x. size (snd x)"}\newline -*** Result matrix:\newline -*** \ \ \ \ 1\ \ 2 \newline -*** a: ? <= \newline -*** Could not find lexicographic termination order.\newline -*** At command "fun".\newline -\end{isabelle} -*} - - -text {* - The key to this error message is the matrix at the bottom. The rows - of that matrix correspond to the different recursive calls (In our - case, there is just one). The columns are the function's arguments - (expressed through different measure functions, which map the - argument tuple to a natural number). - - The contents of the matrix summarize what is known about argument - descents: The second argument has a weak descent (@{text "<="}) at the - recursive call, and for the first argument nothing could be proved, - which is expressed by @{text "?"}. In general, there are the values - @{text "<"}, @{text "<="} and @{text "?"}. - - For the failed proof attempts, the unfinished subgoals are also - printed. Looking at these will often point to a missing lemma. - -% As a more real example, here is quicksort: -*} -(* -function qs :: "nat list \ nat list" -where - "qs [] = []" -| "qs (x#xs) = qs [y\xs. y < x] @ x # qs [y\xs. y \ x]" -by pat_completeness auto - -termination apply lexicographic_order - -text {* If we try @{text "lexicographic_order"} method, we get the - following error *} -termination by (lexicographic_order simp:l2) - -lemma l: "x \ y \ x < Suc y" by arith - -function - -*) - -section {* Mutual Recursion *} - -text {* - If two or more functions call one another mutually, they have to be defined - in one step. Here are @{text "even"} and @{text "odd"}: -*} - -function even :: "nat \ bool" - and odd :: "nat \ bool" -where - "even 0 = True" -| "odd 0 = False" -| "even (Suc n) = odd n" -| "odd (Suc n) = even n" -by pat_completeness auto - -text {* - To eliminate the mutual dependencies, Isabelle internally - creates a single function operating on the sum - type @{typ "nat + nat"}. Then, @{const even} and @{const odd} are - defined as projections. Consequently, termination has to be proved - simultaneously for both functions, by specifying a measure on the - sum type: -*} - -termination -by (relation "measure (\x. case x of Inl n \ n | Inr n \ n)") auto - -text {* - We could also have used @{text lexicographic_order}, which - supports mutual recursive termination proofs to a certain extent. -*} - -subsection {* Induction for mutual recursion *} - -text {* - - When functions are mutually recursive, proving properties about them - generally requires simultaneous induction. The induction rule @{text "even_odd.induct"} - generated from the above definition reflects this. - - Let us prove something about @{const even} and @{const odd}: -*} - -lemma even_odd_mod2: - "even n = (n mod 2 = 0)" - "odd n = (n mod 2 = 1)" - -txt {* - We apply simultaneous induction, specifying the induction variable - for both goals, separated by \cmd{and}: *} - -apply (induct n and n rule: even_odd.induct) - -txt {* - We get four subgoals, which correspond to the clauses in the - definition of @{const even} and @{const odd}: - @{subgoals[display,indent=0]} - Simplification solves the first two goals, leaving us with two - statements about the @{text "mod"} operation to prove: -*} - -apply simp_all - -txt {* - @{subgoals[display,indent=0]} - - \noindent These can be handled by Isabelle's arithmetic decision procedures. - -*} - -apply arith -apply arith -done - -text {* - In proofs like this, the simultaneous induction is really essential: - Even if we are just interested in one of the results, the other - one is necessary to strengthen the induction hypothesis. If we leave - out the statement about @{const odd} and just write @{term True} instead, - the same proof fails: -*} - -lemma failed_attempt: - "even n = (n mod 2 = 0)" - "True" -apply (induct n rule: even_odd.induct) - -txt {* - \noindent Now the third subgoal is a dead end, since we have no - useful induction hypothesis available: - - @{subgoals[display,indent=0]} -*} - -oops - -section {* General pattern matching *} -text{*\label{genpats} *} - -subsection {* Avoiding automatic pattern splitting *} - -text {* - - Up to now, we used pattern matching only on datatypes, and the - patterns were always disjoint and complete, and if they weren't, - they were made disjoint automatically like in the definition of - @{const "sep"} in \S\ref{patmatch}. - - This automatic splitting can significantly increase the number of - equations involved, and this is not always desirable. The following - example shows the problem: - - Suppose we are modeling incomplete knowledge about the world by a - three-valued datatype, which has values @{term "T"}, @{term "F"} - and @{term "X"} for true, false and uncertain propositions, respectively. -*} - -datatype P3 = T | F | X - -text {* \noindent Then the conjunction of such values can be defined as follows: *} - -fun And :: "P3 \ P3 \ P3" -where - "And T p = p" -| "And p T = p" -| "And p F = F" -| "And F p = F" -| "And X X = X" - - -text {* - This definition is useful, because the equations can directly be used - as simplification rules. But the patterns overlap: For example, - the expression @{term "And T T"} is matched by both the first and - the second equation. By default, Isabelle makes the patterns disjoint by - splitting them up, producing instances: -*} - -thm And.simps - -text {* - @{thm[indent=4] And.simps} - - \vspace*{1em} - \noindent There are several problems with this: - - \begin{enumerate} - \item If the datatype has many constructors, there can be an - explosion of equations. For @{const "And"}, we get seven instead of - five equations, which can be tolerated, but this is just a small - example. - - \item Since splitting makes the equations \qt{less general}, they - do not always match in rewriting. While the term @{term "And x F"} - can be simplified to @{term "F"} with the original equations, a - (manual) case split on @{term "x"} is now necessary. - - \item The splitting also concerns the induction rule @{text - "And.induct"}. Instead of five premises it now has seven, which - means that our induction proofs will have more cases. - - \item In general, it increases clarity if we get the same definition - back which we put in. - \end{enumerate} - - If we do not want the automatic splitting, we can switch it off by - leaving out the \cmd{sequential} option. However, we will have to - prove that our pattern matching is consistent\footnote{This prevents - us from defining something like @{term "f x = True"} and @{term "f x - = False"} simultaneously.}: -*} - -function And2 :: "P3 \ P3 \ P3" -where - "And2 T p = p" -| "And2 p T = p" -| "And2 p F = F" -| "And2 F p = F" -| "And2 X X = X" - -txt {* - \noindent Now let's look at the proof obligations generated by a - function definition. In this case, they are: - - @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} - - The first subgoal expresses the completeness of the patterns. It has - the form of an elimination rule and states that every @{term x} of - the function's input type must match at least one of the patterns\footnote{Completeness could - be equivalently stated as a disjunction of existential statements: -@{term "(\p. x = (T, p)) \ (\p. x = (p, T)) \ (\p. x = (p, F)) \ - (\p. x = (F, p)) \ (x = (X, X))"}, and you can use the method @{text atomize_elim} to get that form instead.}. If the patterns just involve - datatypes, we can solve it with the @{text "pat_completeness"} - method: -*} - -apply pat_completeness - -txt {* - The remaining subgoals express \emph{pattern compatibility}. We do - allow that an input value matches multiple patterns, but in this - case, the result (i.e.~the right hand sides of the equations) must - also be equal. For each pair of two patterns, there is one such - subgoal. Usually this needs injectivity of the constructors, which - is used automatically by @{text "auto"}. -*} - -by auto - - -subsection {* Non-constructor patterns *} - -text {* - Most of Isabelle's basic types take the form of inductive datatypes, - and usually pattern matching works on the constructors of such types. - However, this need not be always the case, and the \cmd{function} - command handles other kind of patterns, too. - - One well-known instance of non-constructor patterns are - so-called \emph{$n+k$-patterns}, which are a little controversial in - the functional programming world. Here is the initial fibonacci - example with $n+k$-patterns: -*} - -function fib2 :: "nat \ nat" -where - "fib2 0 = 1" -| "fib2 1 = 1" -| "fib2 (n + 2) = fib2 n + fib2 (Suc n)" - -(*<*)ML_val "goals_limit := 1"(*>*) -txt {* - This kind of matching is again justified by the proof of pattern - completeness and compatibility. - The proof obligation for pattern completeness states that every natural number is - either @{term "0::nat"}, @{term "1::nat"} or @{term "n + - (2::nat)"}: - - @{subgoals[display,indent=0]} - - This is an arithmetic triviality, but unfortunately the - @{text arith} method cannot handle this specific form of an - elimination rule. However, we can use the method @{text - "atomize_elim"} to do an ad-hoc conversion to a disjunction of - existentials, which can then be solved by the arithmetic decision procedure. - Pattern compatibility and termination are automatic as usual. -*} -(*<*)ML_val "goals_limit := 10"(*>*) -apply atomize_elim -apply arith -apply auto -done -termination by lexicographic_order -text {* - We can stretch the notion of pattern matching even more. The - following function is not a sensible functional program, but a - perfectly valid mathematical definition: -*} - -function ev :: "nat \ bool" -where - "ev (2 * n) = True" -| "ev (2 * n + 1) = False" -apply atomize_elim -by arith+ -termination by (relation "{}") simp - -text {* - This general notion of pattern matching gives you a certain freedom - in writing down specifications. However, as always, such freedom should - be used with care: - - If we leave the area of constructor - patterns, we have effectively departed from the world of functional - programming. This means that it is no longer possible to use the - code generator, and expect it to generate ML code for our - definitions. Also, such a specification might not work very well together with - simplification. Your mileage may vary. -*} - - -subsection {* Conditional equations *} - -text {* - The function package also supports conditional equations, which are - similar to guards in a language like Haskell. Here is Euclid's - algorithm written with conditional patterns\footnote{Note that the - patterns are also overlapping in the base case}: -*} - -function gcd :: "nat \ nat \ nat" -where - "gcd x 0 = x" -| "gcd 0 y = y" -| "x < y \ gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)" -| "\ x < y \ gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)" -by (atomize_elim, auto, arith) -termination by lexicographic_order - -text {* - By now, you can probably guess what the proof obligations for the - pattern completeness and compatibility look like. - - Again, functions with conditional patterns are not supported by the - code generator. -*} - - -subsection {* Pattern matching on strings *} - -text {* - As strings (as lists of characters) are normal datatypes, pattern - matching on them is possible, but somewhat problematic. Consider the - following definition: - -\end{isamarkuptext} -\noindent\cmd{fun} @{text "check :: \"string \ bool\""}\\% -\cmd{where}\\% -\hspace*{2ex}@{text "\"check (''good'') = True\""}\\% -@{text "| \"check s = False\""} -\begin{isamarkuptext} - - \noindent An invocation of the above \cmd{fun} command does not - terminate. What is the problem? Strings are lists of characters, and - characters are a datatype with a lot of constructors. Splitting the - catch-all pattern thus leads to an explosion of cases, which cannot - be handled by Isabelle. - - There are two things we can do here. Either we write an explicit - @{text "if"} on the right hand side, or we can use conditional patterns: -*} - -function check :: "string \ bool" -where - "check (''good'') = True" -| "s \ ''good'' \ check s = False" -by auto - - -section {* Partiality *} - -text {* - In HOL, all functions are total. A function @{term "f"} applied to - @{term "x"} always has the value @{term "f x"}, and there is no notion - of undefinedness. - This is why we have to do termination - proofs when defining functions: The proof justifies that the - function can be defined by wellfounded recursion. - - However, the \cmd{function} package does support partiality to a - certain extent. Let's look at the following function which looks - for a zero of a given function f. -*} - -function (*<*)(domintros, tailrec)(*>*)findzero :: "(nat \ nat) \ nat \ nat" -where - "findzero f n = (if f n = 0 then n else findzero f (Suc n))" -by pat_completeness auto -(*<*)declare findzero.simps[simp del](*>*) - -text {* - \noindent Clearly, any attempt of a termination proof must fail. And without - that, we do not get the usual rules @{text "findzero.simps"} and - @{text "findzero.induct"}. So what was the definition good for at all? -*} - -subsection {* Domain predicates *} - -text {* - The trick is that Isabelle has not only defined the function @{const findzero}, but also - a predicate @{term "findzero_dom"} that characterizes the values where the function - terminates: the \emph{domain} of the function. If we treat a - partial function just as a total function with an additional domain - predicate, we can derive simplification and - induction rules as we do for total functions. They are guarded - by domain conditions and are called @{text psimps} and @{text - pinduct}: -*} - -text {* - \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage} - \hfill(@{text "findzero.psimps"}) - \vspace{1em} - - \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage} - \hfill(@{text "findzero.pinduct"}) -*} - -text {* - Remember that all we - are doing here is use some tricks to make a total function appear - as if it was partial. We can still write the term @{term "findzero - (\x. 1) 0"} and like any other term of type @{typ nat} it is equal - to some natural number, although we might not be able to find out - which one. The function is \emph{underdefined}. - - But it is defined enough to prove something interesting about it. We - can prove that if @{term "findzero f n"} - terminates, it indeed returns a zero of @{term f}: -*} - -lemma findzero_zero: "findzero_dom (f, n) \ f (findzero f n) = 0" - -txt {* \noindent We apply induction as usual, but using the partial induction - rule: *} - -apply (induct f n rule: findzero.pinduct) - -txt {* \noindent This gives the following subgoals: - - @{subgoals[display,indent=0]} - - \noindent The hypothesis in our lemma was used to satisfy the first premise in - the induction rule. However, we also get @{term - "findzero_dom (f, n)"} as a local assumption in the induction step. This - allows to unfold @{term "findzero f n"} using the @{text psimps} - rule, and the rest is trivial. Since the @{text psimps} rules carry the - @{text "[simp]"} attribute by default, we just need a single step: - *} -apply simp -done - -text {* - Proofs about partial functions are often not harder than for total - functions. Fig.~\ref{findzero_isar} shows a slightly more - complicated proof written in Isar. It is verbose enough to show how - partiality comes into play: From the partial induction, we get an - additional domain condition hypothesis. Observe how this condition - is applied when calls to @{term findzero} are unfolded. -*} - -text_raw {* -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -*} -lemma "\findzero_dom (f, n); x \ {n ..< findzero f n}\ \ f x \ 0" -proof (induct rule: findzero.pinduct) - fix f n assume dom: "findzero_dom (f, n)" - and IH: "\f n \ 0; x \ {Suc n ..< findzero f (Suc n)}\ \ f x \ 0" - and x_range: "x \ {n ..< findzero f n}" - have "f n \ 0" - proof - assume "f n = 0" - with dom have "findzero f n = n" by simp - with x_range show False by auto - qed - - from x_range have "x = n \ x \ {Suc n ..< findzero f n}" by auto - thus "f x \ 0" - proof - assume "x = n" - with `f n \ 0` show ?thesis by simp - next - assume "x \ {Suc n ..< findzero f n}" - with dom and `f n \ 0` have "x \ {Suc n ..< findzero f (Suc n)}" by simp - with IH and `f n \ 0` - show ?thesis by simp - qed -qed -text_raw {* -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{A proof about a partial function}\label{findzero_isar} -\end{figure} -*} - -subsection {* Partial termination proofs *} - -text {* - Now that we have proved some interesting properties about our - function, we should turn to the domain predicate and see if it is - actually true for some values. Otherwise we would have just proved - lemmas with @{term False} as a premise. - - Essentially, we need some introduction rules for @{text - findzero_dom}. The function package can prove such domain - introduction rules automatically. But since they are not used very - often (they are almost never needed if the function is total), this - functionality is disabled by default for efficiency reasons. So we have to go - back and ask for them explicitly by passing the @{text - "(domintros)"} option to the function package: - -\vspace{1ex} -\noindent\cmd{function} @{text "(domintros) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% -\cmd{where}\isanewline% -\ \ \ldots\\ - - \noindent Now the package has proved an introduction rule for @{text findzero_dom}: -*} - -thm findzero.domintros - -text {* - @{thm[display] findzero.domintros} - - Domain introduction rules allow to show that a given value lies in the - domain of a function, if the arguments of all recursive calls - are in the domain as well. They allow to do a \qt{single step} in a - termination proof. Usually, you want to combine them with a suitable - induction principle. - - Since our function increases its argument at recursive calls, we - need an induction principle which works \qt{backwards}. We will use - @{text inc_induct}, which allows to do induction from a fixed number - \qt{downwards}: - - \begin{center}@{thm inc_induct}\hfill(@{text "inc_induct"})\end{center} - - Figure \ref{findzero_term} gives a detailed Isar proof of the fact - that @{text findzero} terminates if there is a zero which is greater - or equal to @{term n}. First we derive two useful rules which will - solve the base case and the step case of the induction. The - induction is then straightforward, except for the unusual induction - principle. - -*} - -text_raw {* -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -*} -lemma findzero_termination: - assumes "x \ n" and "f x = 0" - shows "findzero_dom (f, n)" -proof - - have base: "findzero_dom (f, x)" - by (rule findzero.domintros) (simp add:`f x = 0`) - - have step: "\i. findzero_dom (f, Suc i) - \ findzero_dom (f, i)" - by (rule findzero.domintros) simp - - from `x \ n` show ?thesis - proof (induct rule:inc_induct) - show "findzero_dom (f, x)" by (rule base) - next - fix i assume "findzero_dom (f, Suc i)" - thus "findzero_dom (f, i)" by (rule step) - qed -qed -text_raw {* -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{Termination proof for @{text findzero}}\label{findzero_term} -\end{figure} -*} - -text {* - Again, the proof given in Fig.~\ref{findzero_term} has a lot of - detail in order to explain the principles. Using more automation, we - can also have a short proof: -*} - -lemma findzero_termination_short: - assumes zero: "x >= n" - assumes [simp]: "f x = 0" - shows "findzero_dom (f, n)" -using zero -by (induct rule:inc_induct) (auto intro: findzero.domintros) - -text {* - \noindent It is simple to combine the partial correctness result with the - termination lemma: -*} - -lemma findzero_total_correctness: - "f x = 0 \ f (findzero f 0) = 0" -by (blast intro: findzero_zero findzero_termination) - -subsection {* Definition of the domain predicate *} - -text {* - Sometimes it is useful to know what the definition of the domain - predicate looks like. Actually, @{text findzero_dom} is just an - abbreviation: - - @{abbrev[display] findzero_dom} - - The domain predicate is the \emph{accessible part} of a relation @{const - findzero_rel}, which was also created internally by the function - package. @{const findzero_rel} is just a normal - inductive predicate, so we can inspect its definition by - looking at the introduction rules @{text findzero_rel.intros}. - In our case there is just a single rule: - - @{thm[display] findzero_rel.intros} - - The predicate @{const findzero_rel} - describes the \emph{recursion relation} of the function - definition. The recursion relation is a binary relation on - the arguments of the function that relates each argument to its - recursive calls. In general, there is one introduction rule for each - recursive call. - - The predicate @{term "accp findzero_rel"} is the accessible part of - that relation. An argument belongs to the accessible part, if it can - be reached in a finite number of steps (cf.~its definition in @{text - "Wellfounded.thy"}). - - Since the domain predicate is just an abbreviation, you can use - lemmas for @{const accp} and @{const findzero_rel} directly. Some - lemmas which are occasionally useful are @{text accpI}, @{text - accp_downward}, and of course the introduction and elimination rules - for the recursion relation @{text "findzero.intros"} and @{text "findzero.cases"}. -*} - -(*lemma findzero_nicer_domintros: - "f x = 0 \ findzero_dom (f, x)" - "findzero_dom (f, Suc x) \ findzero_dom (f, x)" -by (rule accpI, erule findzero_rel.cases, auto)+ -*) - -subsection {* A Useful Special Case: Tail recursion *} - -text {* - The domain predicate is our trick that allows us to model partiality - in a world of total functions. The downside of this is that we have - to carry it around all the time. The termination proof above allowed - us to replace the abstract @{term "findzero_dom (f, n)"} by the more - concrete @{term "(x \ n \ f x = (0::nat))"}, but the condition is still - there and can only be discharged for special cases. - In particular, the domain predicate guards the unfolding of our - function, since it is there as a condition in the @{text psimp} - rules. - - Now there is an important special case: We can actually get rid - of the condition in the simplification rules, \emph{if the function - is tail-recursive}. The reason is that for all tail-recursive - equations there is a total function satisfying them, even if they - are non-terminating. - -% A function is tail recursive, if each call to the function is either -% equal -% -% So the outer form of the -% -%if it can be written in the following -% form: -% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} - - - The function package internally does the right construction and can - derive the unconditional simp rules, if we ask it to do so. Luckily, - our @{const "findzero"} function is tail-recursive, so we can just go - back and add another option to the \cmd{function} command: - -\vspace{1ex} -\noindent\cmd{function} @{text "(domintros, tailrec) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% -\cmd{where}\isanewline% -\ \ \ldots\\% - - - \noindent Now, we actually get unconditional simplification rules, even - though the function is partial: -*} - -thm findzero.simps - -text {* - @{thm[display] findzero.simps} - - \noindent Of course these would make the simplifier loop, so we better remove - them from the simpset: -*} - -declare findzero.simps[simp del] - -text {* - Getting rid of the domain conditions in the simplification rules is - not only useful because it simplifies proofs. It is also required in - order to use Isabelle's code generator to generate ML code - from a function definition. - Since the code generator only works with equations, it cannot be - used with @{text "psimp"} rules. Thus, in order to generate code for - partial functions, they must be defined as a tail recursion. - Luckily, many functions have a relatively natural tail recursive - definition. -*} - -section {* Nested recursion *} - -text {* - Recursive calls which are nested in one another frequently cause - complications, since their termination proof can depend on a partial - correctness property of the function itself. - - As a small example, we define the \qt{nested zero} function: -*} - -function nz :: "nat \ nat" -where - "nz 0 = 0" -| "nz (Suc n) = nz (nz n)" -by pat_completeness auto - -text {* - If we attempt to prove termination using the identity measure on - naturals, this fails: -*} - -termination - apply (relation "measure (\n. n)") - apply auto - -txt {* - We get stuck with the subgoal - - @{subgoals[display]} - - Of course this statement is true, since we know that @{const nz} is - the zero function. And in fact we have no problem proving this - property by induction. -*} -(*<*)oops(*>*) -lemma nz_is_zero: "nz_dom n \ nz n = 0" - by (induct rule:nz.pinduct) auto - -text {* - We formulate this as a partial correctness lemma with the condition - @{term "nz_dom n"}. This allows us to prove it with the @{text - pinduct} rule before we have proved termination. With this lemma, - the termination proof works as expected: -*} - -termination - by (relation "measure (\n. n)") (auto simp: nz_is_zero) - -text {* - As a general strategy, one should prove the statements needed for - termination as a partial property first. Then they can be used to do - the termination proof. This also works for less trivial - examples. Figure \ref{f91} defines the 91-function, a well-known - challenge problem due to John McCarthy, and proves its termination. -*} - -text_raw {* -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -*} - -function f91 :: "nat \ nat" -where - "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" -by pat_completeness auto - -lemma f91_estimate: - assumes trm: "f91_dom n" - shows "n < f91 n + 11" -using trm by induct auto - -termination -proof - let ?R = "measure (\x. 101 - x)" - show "wf ?R" .. - - fix n :: nat assume "\ 100 < n" -- "Assumptions for both calls" - - thus "(n + 11, n) \ ?R" by simp -- "Inner call" - - assume inner_trm: "f91_dom (n + 11)" -- "Outer call" - with f91_estimate have "n + 11 < f91 (n + 11) + 11" . - with `\ 100 < n` show "(f91 (n + 11), n) \ ?R" by simp -qed - -text_raw {* -\isamarkupfalse\isabellestyle{tt} -\end{minipage} -\vspace{6pt}\hrule -\caption{McCarthy's 91-function}\label{f91} -\end{figure} -*} - - -section {* Higher-Order Recursion *} - -text {* - Higher-order recursion occurs when recursive calls - are passed as arguments to higher-order combinators such as @{const - map}, @{term filter} etc. - As an example, imagine a datatype of n-ary trees: -*} - -datatype 'a tree = - Leaf 'a -| Branch "'a tree list" - - -text {* \noindent We can define a function which swaps the left and right subtrees recursively, using the - list functions @{const rev} and @{const map}: *} - -fun mirror :: "'a tree \ 'a tree" -where - "mirror (Leaf n) = Leaf n" -| "mirror (Branch l) = Branch (rev (map mirror l))" - -text {* - Although the definition is accepted without problems, let us look at the termination proof: -*} - -termination proof - txt {* - - As usual, we have to give a wellfounded relation, such that the - arguments of the recursive calls get smaller. But what exactly are - the arguments of the recursive calls when mirror is given as an - argument to @{const map}? Isabelle gives us the - subgoals - - @{subgoals[display,indent=0]} - - So the system seems to know that @{const map} only - applies the recursive call @{term "mirror"} to elements - of @{term "l"}, which is essential for the termination proof. - - This knowledge about @{const map} is encoded in so-called congruence rules, - which are special theorems known to the \cmd{function} command. The - rule for @{const map} is - - @{thm[display] map_cong} - - You can read this in the following way: Two applications of @{const - map} are equal, if the list arguments are equal and the functions - coincide on the elements of the list. This means that for the value - @{term "map f l"} we only have to know how @{term f} behaves on - the elements of @{term l}. - - Usually, one such congruence rule is - needed for each higher-order construct that is used when defining - new functions. In fact, even basic functions like @{const - If} and @{const Let} are handled by this mechanism. The congruence - rule for @{const If} states that the @{text then} branch is only - relevant if the condition is true, and the @{text else} branch only if it - is false: - - @{thm[display] if_cong} - - Congruence rules can be added to the - function package by giving them the @{term fundef_cong} attribute. - - The constructs that are predefined in Isabelle, usually - come with the respective congruence rules. - But if you define your own higher-order functions, you may have to - state and prove the required congruence rules yourself, if you want to use your - functions in recursive definitions. -*} -(*<*)oops(*>*) - -subsection {* Congruence Rules and Evaluation Order *} - -text {* - Higher order logic differs from functional programming languages in - that it has no built-in notion of evaluation order. A program is - just a set of equations, and it is not specified how they must be - evaluated. - - However for the purpose of function definition, we must talk about - evaluation order implicitly, when we reason about termination. - Congruence rules express that a certain evaluation order is - consistent with the logical definition. - - Consider the following function. -*} - -function f :: "nat \ bool" -where - "f n = (n = 0 \ f (n - 1))" -(*<*)by pat_completeness auto(*>*) - -text {* - For this definition, the termination proof fails. The default configuration - specifies no congruence rule for disjunction. We have to add a - congruence rule that specifies left-to-right evaluation order: - - \vspace{1ex} - \noindent @{thm disj_cong}\hfill(@{text "disj_cong"}) - \vspace{1ex} - - Now the definition works without problems. Note how the termination - proof depends on the extra condition that we get from the congruence - rule. - - However, as evaluation is not a hard-wired concept, we - could just turn everything around by declaring a different - congruence rule. Then we can make the reverse definition: -*} - -lemma disj_cong2[fundef_cong]: - "(\ Q' \ P = P') \ (Q = Q') \ (P \ Q) = (P' \ Q')" - by blast - -fun f' :: "nat \ bool" -where - "f' n = (f' (n - 1) \ n = 0)" - -text {* - \noindent These examples show that, in general, there is no \qt{best} set of - congruence rules. - - However, such tweaking should rarely be necessary in - practice, as most of the time, the default set of congruence rules - works well. -*} - -end diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/Thy/ROOT.ML --- a/doc-src/IsarAdvanced/Functions/Thy/ROOT.ML Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,4 +0,0 @@ - -(* $Id$ *) - -use_thy "Functions"; diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex --- a/doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1985 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Functions}% -% -\isadelimtheory -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Functions\isanewline -\isakeyword{imports}\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{Function Definitions for Dummies% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In most cases, defining a recursive function is just as simple as other definitions:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ n\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The syntax is rather self-explanatory: We introduce a function by - giving its name, its type, - and a set of defining recursive equations. - If we leave out the type, the most general type will be - inferred, which can sometimes lead to surprises: Since both \isa{{\isadigit{1}}} and \isa{{\isacharplus}} are overloaded, we would end up - with \isa{fib\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isacharcolon}{\isacharcolon}{\isacharbraceleft}one{\isacharcomma}plus{\isacharbraceright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The function always terminates, since its argument gets smaller in - every recursive call. - Since HOL is a logic of total functions, termination is a - fundamental requirement to prevent inconsistencies\footnote{From the - \qt{definition} \isa{f{\isacharparenleft}n{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}n{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}} we could prove - \isa{{\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}} by subtracting \isa{f{\isacharparenleft}n{\isacharparenright}} on both sides.}. - Isabelle tries to prove termination automatically when a definition - is made. In \S\ref{termination}, we will look at cases where this - fails and see what to do then.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Pattern matching% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{patmatch} - Like in functional programming, we can use pattern matching to - define functions. At the moment we will only consider \emph{constructor - patterns}, which only consist of datatype constructors and - variables. Furthermore, patterns must be linear, i.e.\ all variables - on the left hand side of an equation must be distinct. In - \S\ref{genpats} we discuss more general pattern matching. - - If patterns overlap, the order of the equations is taken into - account. The following function inserts a fixed element between any - two elements of a list:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}sep\ a\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}sep\ a\ xs\ \ \ \ \ \ \ {\isacharequal}\ xs{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Overlapping patterns are interpreted as \qt{increments} to what is - already there: The second equation is only meant for the cases where - the first one does not match. Consequently, Isabelle replaces it - internally by the remaining cases, making the patterns disjoint:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ sep{\isachardot}simps% -\begin{isamarkuptext}% -\begin{isabelle}% -sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\isasep\isanewline% -sep\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\isasep\isanewline% -sep\ a\ {\isacharbrackleft}v{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}v{\isacharbrackright}% -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent The equations from function definitions are automatically used in - simplification:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}sep\ {\isadigit{0}}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Induction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle provides customized induction rules for recursive - functions. These rules follow the recursive structure of the - definition. Here is the rule \isa{sep{\isachardot}induct} arising from the - above definition of \isa{sep}: - - \begin{isabelle}% -{\isasymlbrakk}{\isasymAnd}a\ x\ y\ xs{\isachardot}\ {\isacharquery}P\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}P\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharsemicolon}\ {\isasymAnd}a{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ {\isasymAnd}a\ v{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isasymrbrakk}\isanewline -{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% -\end{isabelle} - - We have a step case for list with at least two elements, and two - base cases for the zero- and the one-element list. Here is a simple - proof about \isa{sep} and \isa{map}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}map\ f\ {\isacharparenleft}sep\ x\ ys{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}map\ f\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ x\ ys\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}% -\begin{isamarkuptxt}% -We get three cases, like in the definition. - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ x\ y\ xs{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ v{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ auto\ \isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -With the \cmd{fun} command, you can define about 80\% of the - functions that occur in practice. The rest of this tutorial explains - the remaining 20\%.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{fun vs.\ function% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \cmd{fun} command provides a - convenient shorthand notation for simple function definitions. In - this mode, Isabelle tries to solve all the necessary proof obligations - automatically. If any proof fails, the definition is - rejected. This can either mean that the definition is indeed faulty, - or that the default proof procedures are just not smart enough (or - rather: not designed) to handle the definition. - - By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or - solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: - -\end{isamarkuptext} - - -\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} -\cmd{fun} \isa{f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\vspace*{6pt} -\end{minipage}\right] -\quad\equiv\quad -\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} -\cmd{function} \isa{{\isacharparenleft}}\cmd{sequential}\isa{{\isacharparenright}\ f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\\% -\cmd{by} \isa{pat{\isacharunderscore}completeness\ auto}\\% -\cmd{termination by} \isa{lexicographic{\isacharunderscore}order}\vspace{6pt} -\end{minipage} -\right]\] - -\begin{isamarkuptext} - \vspace*{1em} - \noindent Some details have now become explicit: - - \begin{enumerate} - \item The \cmd{sequential} option enables the preprocessing of - pattern overlaps which we already saw. Without this option, the equations - must already be disjoint and complete. The automatic completion only - works with constructor patterns. - - \item A function definition produces a proof obligation which - expresses completeness and compatibility of patterns (we talk about - this later). The combination of the methods \isa{pat{\isacharunderscore}completeness} and - \isa{auto} is used to solve this proof obligation. - - \item A termination proof follows the definition, started by the - \cmd{termination} command. This will be explained in \S\ref{termination}. - \end{enumerate} - Whenever a \cmd{fun} command fails, it is usually a good idea to - expand the syntax to the more verbose \cmd{function} form, to see - what is actually going on.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Termination% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{termination} - The method \isa{lexicographic{\isacharunderscore}order} is the default method for - termination proofs. It can prove termination of a - certain class of functions by searching for a suitable lexicographic - combination of size measures. Of course, not all functions have such - a simple termination argument. For them, we can specify the termination - relation manually.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The {\tt relation} method% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function, which sums up natural numbers up to - \isa{N}, using a counter \isa{i}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ sum\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}sum\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ then\ {\isadigit{0}}\ else\ i\ {\isacharplus}\ sum\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent The \isa{lexicographic{\isacharunderscore}order} method fails on this example, because none of the - arguments decreases in the recursive call, with respect to the standard size ordering. - To prove termination manually, we must provide a custom wellfounded relation. - - The termination argument for \isa{sum} is based on the fact that - the \emph{difference} between \isa{i} and \isa{N} gets - smaller in every step, and that the recursion stops when \isa{i} - is greater than \isa{N}. Phrased differently, the expression - \isa{N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i} always decreases. - - We can use this expression as a measure function suitable to prove termination.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\ sum\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\begin{isamarkuptxt}% -The \cmd{termination} command sets up the termination goal for the - specified function \isa{sum}. If the function name is omitted, it - implicitly refers to the last function definition. - - The \isa{relation} method takes a relation of - type \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}, where \isa{{\isacharprime}a} is the argument type of - the function. If the function has multiple curried arguments, then - these are packed together into a tuple, as it happened in the above - example. - - The predefined function \isa{{\isachardoublequote}measure\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}} constructs a - wellfounded relation from a mapping into the natural numbers (a - \emph{measure function}). - - After the invocation of \isa{relation}, we must prove that (a) - the relation we supplied is wellfounded, and (b) that the arguments - of recursive calls indeed decrease with respect to the - relation: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharparenleft}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}i\ N{\isachardot}\ {\isasymnot}\ N\ {\isacharless}\ i\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharparenleft}Suc\ i{\isacharcomma}\ N{\isacharparenright}{\isacharcomma}\ i{\isacharcomma}\ N{\isacharparenright}\ {\isasymin}\ measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}% -\end{isabelle} - - These goals are all solved by \isa{auto}:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ auto\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Let us complicate the function a little, by adding some more - recursive calls:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ foo\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}foo\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ \isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ then\ {\isacharparenleft}if\ N\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isadigit{0}}\ else\ foo\ {\isadigit{0}}\ {\isacharparenleft}N\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ i\ {\isacharplus}\ foo\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -When \isa{i} has reached \isa{N}, it starts at zero again - and \isa{N} is decremented. - This corresponds to a nested - loop where one index counts up and the other down. Termination can - be proved using a lexicographic combination of two measures, namely - the value of \isa{N} and the above difference. The \isa{measures} combinator generalizes \isa{measure} by taking a - list of measure functions.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\ \isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measures\ {\isacharbrackleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N{\isacharcomma}\ {\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{How \isa{lexicographic{\isacharunderscore}order} works% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -To see how the automatic termination proofs work, let's look at an - example where it fails\footnote{For a detailed discussion of the - termination prover, see \cite{bulwahnKN07}}: - -\end{isamarkuptext} -\cmd{fun} \isa{fails\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat\ list\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% -\cmd{where}\\% -\hspace*{2ex}\isa{{\isachardoublequote}fails\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ a{\isachardoublequote}}\\% -|\hspace*{1.5ex}\isa{{\isachardoublequote}fails\ a\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ fails\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharparenright}\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}{\isachardoublequote}}\\ -\begin{isamarkuptext} - -\noindent Isabelle responds with the following error: - -\begin{isabelle} -*** Unfinished subgoals:\newline -*** (a, 1, <):\newline -*** \ 1.~\isa{{\isasymAnd}x{\isachardot}\ x\ {\isacharequal}\ {\isadigit{0}}}\newline -*** (a, 1, <=):\newline -*** \ 1.~False\newline -*** (a, 2, <):\newline -*** \ 1.~False\newline -*** Calls:\newline -*** a) \isa{{\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharminus}{\isacharminus}{\isachargreater}{\isachargreater}\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}}\newline -*** Measures:\newline -*** 1) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}fst\ x{\isacharparenright}}\newline -*** 2) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}snd\ x{\isacharparenright}}\newline -*** Result matrix:\newline -*** \ \ \ \ 1\ \ 2 \newline -*** a: ? <= \newline -*** Could not find lexicographic termination order.\newline -*** At command "fun".\newline -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The key to this error message is the matrix at the bottom. The rows - of that matrix correspond to the different recursive calls (In our - case, there is just one). The columns are the function's arguments - (expressed through different measure functions, which map the - argument tuple to a natural number). - - The contents of the matrix summarize what is known about argument - descents: The second argument has a weak descent (\isa{{\isacharless}{\isacharequal}}) at the - recursive call, and for the first argument nothing could be proved, - which is expressed by \isa{{\isacharquery}}. In general, there are the values - \isa{{\isacharless}}, \isa{{\isacharless}{\isacharequal}} and \isa{{\isacharquery}}. - - For the failed proof attempts, the unfinished subgoals are also - printed. Looking at these will often point to a missing lemma. - -% As a more real example, here is quicksort:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Mutual Recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -If two or more functions call one another mutually, they have to be defined - in one step. Here are \isa{even} and \isa{odd}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ even\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ odd\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}even\ {\isadigit{0}}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isadigit{0}}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ odd\ n{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ even\ n{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -To eliminate the mutual dependencies, Isabelle internally - creates a single function operating on the sum - type \isa{nat\ {\isacharplus}\ nat}. Then, \isa{even} and \isa{odd} are - defined as projections. Consequently, termination has to be proved - simultaneously for both functions, by specifying a measure on the - sum type:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\ \isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ case\ x\ of\ Inl\ n\ {\isasymRightarrow}\ n\ {\isacharbar}\ Inr\ n\ {\isasymRightarrow}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -We could also have used \isa{lexicographic{\isacharunderscore}order}, which - supports mutual recursive termination proofs to a certain extent.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Induction for mutual recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When functions are mutually recursive, proving properties about them - generally requires simultaneous induction. The induction rule \isa{even{\isacharunderscore}odd{\isachardot}induct} - generated from the above definition reflects this. - - Let us prove something about \isa{even} and \isa{odd}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ even{\isacharunderscore}odd{\isacharunderscore}mod{\isadigit{2}}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -We apply simultaneous induction, specifying the induction variable - for both goals, separated by \cmd{and}:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ n\ \isakeyword{and}\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% -\begin{isamarkuptxt}% -We get four subgoals, which correspond to the clauses in the - definition of \isa{even} and \isa{odd}: - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ odd\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}% -\end{isabelle} - Simplification solves the first two goals, leaving us with two - statements about the \isa{mod} operation to prove:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}% -\end{isabelle} - - \noindent These can be handled by Isabelle's arithmetic decision procedures.% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ arith\isanewline -\isacommand{apply}\isamarkupfalse% -\ arith\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -In proofs like this, the simultaneous induction is really essential: - Even if we are just interested in one of the results, the other - one is necessary to strengthen the induction hypothesis. If we leave - out the statement about \isa{odd} and just write \isa{True} instead, - the same proof fails:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ failed{\isacharunderscore}attempt{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}True{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% -\begin{isamarkuptxt}% -\noindent Now the third subgoal is a dead end, since we have no - useful induction hypothesis available: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ True\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ True\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ True% -\end{isabelle}% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{oops}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{General pattern matching% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{genpats}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Avoiding automatic pattern splitting% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Up to now, we used pattern matching only on datatypes, and the - patterns were always disjoint and complete, and if they weren't, - they were made disjoint automatically like in the definition of - \isa{sep} in \S\ref{patmatch}. - - This automatic splitting can significantly increase the number of - equations involved, and this is not always desirable. The following - example shows the problem: - - Suppose we are modeling incomplete knowledge about the world by a - three-valued datatype, which has values \isa{T}, \isa{F} - and \isa{X} for true, false and uncertain propositions, respectively.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ P{\isadigit{3}}\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ X% -\begin{isamarkuptext}% -\noindent Then the conjunction of such values can be defined as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ And\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}And\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% -\begin{isamarkuptext}% -This definition is useful, because the equations can directly be used - as simplification rules. But the patterns overlap: For example, - the expression \isa{And\ T\ T} is matched by both the first and - the second equation. By default, Isabelle makes the patterns disjoint by - splitting them up, producing instances:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ And{\isachardot}simps% -\begin{isamarkuptext}% -\isa{And\ T\ {\isacharquery}p\ {\isacharequal}\ {\isacharquery}p\isasep\isanewline% -And\ F\ T\ {\isacharequal}\ F\isasep\isanewline% -And\ X\ T\ {\isacharequal}\ X\isasep\isanewline% -And\ F\ F\ {\isacharequal}\ F\isasep\isanewline% -And\ X\ F\ {\isacharequal}\ F\isasep\isanewline% -And\ F\ X\ {\isacharequal}\ F\isasep\isanewline% -And\ X\ X\ {\isacharequal}\ X} - - \vspace*{1em} - \noindent There are several problems with this: - - \begin{enumerate} - \item If the datatype has many constructors, there can be an - explosion of equations. For \isa{And}, we get seven instead of - five equations, which can be tolerated, but this is just a small - example. - - \item Since splitting makes the equations \qt{less general}, they - do not always match in rewriting. While the term \isa{And\ x\ F} - can be simplified to \isa{F} with the original equations, a - (manual) case split on \isa{x} is now necessary. - - \item The splitting also concerns the induction rule \isa{And{\isachardot}induct}. Instead of five premises it now has seven, which - means that our induction proofs will have more cases. - - \item In general, it increases clarity if we get the same definition - back which we put in. - \end{enumerate} - - If we do not want the automatic splitting, we can switch it off by - leaving out the \cmd{sequential} option. However, we will have to - prove that our pattern matching is consistent\footnote{This prevents - us from defining something like \isa{f\ x\ {\isacharequal}\ True} and \isa{f\ x\ {\isacharequal}\ False} simultaneously.}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ And{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}And{\isadigit{2}}\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -\noindent Now let's look at the proof obligations generated by a - function definition. In this case, they are: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ \ }{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ }{\isasymLongrightarrow}\ P\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{5}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X\isanewline -\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline -\ {\isadigit{8}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{9}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{1}}{\isadigit{0}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X% -\end{isabelle}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} - - The first subgoal expresses the completeness of the patterns. It has - the form of an elimination rule and states that every \isa{x} of - the function's input type must match at least one of the patterns\footnote{Completeness could - be equivalently stated as a disjunction of existential statements: -\isa{{\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}}, and you can use the method \isa{atomize{\isacharunderscore}elim} to get that form instead.}. If the patterns just involve - datatypes, we can solve it with the \isa{pat{\isacharunderscore}completeness} - method:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ pat{\isacharunderscore}completeness% -\begin{isamarkuptxt}% -The remaining subgoals express \emph{pattern compatibility}. We do - allow that an input value matches multiple patterns, but in this - case, the result (i.e.~the right hand sides of the equations) must - also be equal. For each pair of two patterns, there is one such - subgoal. Usually this needs injectivity of the constructors, which - is used automatically by \isa{auto}.% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{by}\isamarkupfalse% -\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Non-constructor patterns% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Most of Isabelle's basic types take the form of inductive datatypes, - and usually pattern matching works on the constructors of such types. - However, this need not be always the case, and the \cmd{function} - command handles other kind of patterns, too. - - One well-known instance of non-constructor patterns are - so-called \emph{$n+k$-patterns}, which are a little controversial in - the functional programming world. Here is the initial fibonacci - example with $n+k$-patterns:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ fib{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}\ n\ {\isacharplus}\ fib{\isadigit{2}}\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -This kind of matching is again justified by the proof of pattern - completeness and compatibility. - The proof obligation for pattern completeness states that every natural number is - either \isa{{\isadigit{0}}}, \isa{{\isadigit{1}}} or \isa{n\ {\isacharplus}\ {\isadigit{2}}}: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline -\ {\isadigit{5}}{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline -\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline -\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}n\ na{\isachardot}\isanewline -\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }n\ {\isacharplus}\ {\isadigit{2}}\ {\isacharequal}\ na\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ na\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ na{\isacharparenright}% -\end{isabelle} - - This is an arithmetic triviality, but unfortunately the - \isa{arith} method cannot handle this specific form of an - elimination rule. However, we can use the method \isa{atomize{\isacharunderscore}elim} to do an ad-hoc conversion to a disjunction of - existentials, which can then be solved by the arithmetic decision procedure. - Pattern compatibility and termination are automatic as usual.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ atomize{\isacharunderscore}elim\isanewline -\isacommand{apply}\isamarkupfalse% -\ arith\isanewline -\isacommand{apply}\isamarkupfalse% -\ auto\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ lexicographic{\isacharunderscore}order% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -We can stretch the notion of pattern matching even more. The - following function is not a sensible functional program, but a - perfectly valid mathematical definition:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ ev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ atomize{\isacharunderscore}elim\isanewline -\isacommand{by}\isamarkupfalse% -\ arith{\isacharplus}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}{\isacharparenright}\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -This general notion of pattern matching gives you a certain freedom - in writing down specifications. However, as always, such freedom should - be used with care: - - If we leave the area of constructor - patterns, we have effectively departed from the world of functional - programming. This means that it is no longer possible to use the - code generator, and expect it to generate ML code for our - definitions. Also, such a specification might not work very well together with - simplification. Your mileage may vary.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Conditional equations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The function package also supports conditional equations, which are - similar to guards in a language like Haskell. Here is Euclid's - algorithm written with conditional patterns\footnote{Note that the - patterns are also overlapping in the base case}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}gcd\ x\ {\isadigit{0}}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}gcd\ {\isadigit{0}}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}y\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}{\isasymnot}\ x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}x\ {\isacharminus}\ y{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}atomize{\isacharunderscore}elim{\isacharcomma}\ auto{\isacharcomma}\ arith{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ lexicographic{\isacharunderscore}order% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -By now, you can probably guess what the proof obligations for the - pattern completeness and compatibility look like. - - Again, functions with conditional patterns are not supported by the - code generator.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Pattern matching on strings% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As strings (as lists of characters) are normal datatypes, pattern - matching on them is possible, but somewhat problematic. Consider the - following definition: - -\end{isamarkuptext} -\noindent\cmd{fun} \isa{check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}string\ {\isasymRightarrow}\ bool{\isachardoublequote}}\\% -\cmd{where}\\% -\hspace*{2ex}\isa{{\isachardoublequote}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}}\\% -\isa{{\isacharbar}\ {\isachardoublequote}check\ s\ {\isacharequal}\ False{\isachardoublequote}} -\begin{isamarkuptext} - - \noindent An invocation of the above \cmd{fun} command does not - terminate. What is the problem? Strings are lists of characters, and - characters are a datatype with a lot of constructors. Splitting the - catch-all pattern thus leads to an explosion of cases, which cannot - be handled by Isabelle. - - There are two things we can do here. Either we write an explicit - \isa{if} on the right hand side, or we can use conditional patterns:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}s\ {\isasymnoteq}\ {\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}\ {\isasymLongrightarrow}\ check\ s\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{Partiality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In HOL, all functions are total. A function \isa{f} applied to - \isa{x} always has the value \isa{f\ x}, and there is no notion - of undefinedness. - This is why we have to do termination - proofs when defining functions: The proof justifies that the - function can be defined by wellfounded recursion. - - However, the \cmd{function} package does support partiality to a - certain extent. Let's look at the following function which looks - for a zero of a given function f.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ {\isacharparenleft}if\ f\ n\ {\isacharequal}\ {\isadigit{0}}\ then\ n\ else\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Clearly, any attempt of a termination proof must fail. And without - that, we do not get the usual rules \isa{findzero{\isachardot}simps} and - \isa{findzero{\isachardot}induct}. So what was the definition good for at all?% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Domain predicates% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The trick is that Isabelle has not only defined the function \isa{findzero}, but also - a predicate \isa{findzero{\isacharunderscore}dom} that characterizes the values where the function - terminates: the \emph{domain} of the function. If we treat a - partial function just as a total function with an additional domain - predicate, we can derive simplification and - induction rules as we do for total functions. They are guarded - by domain conditions and are called \isa{psimps} and \isa{pinduct}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% -findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}\ {\isasymLongrightarrow}\isanewline -findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% -\end{isabelle}\end{minipage} - \hfill(\isa{findzero{\isachardot}psimps}) - \vspace{1em} - - \noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% -{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}{\isacharcomma}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}{\isacharparenright}{\isacharsemicolon}\isanewline -\isaindent{\ }{\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ n{\isasymrbrakk}\isanewline -{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% -\end{isabelle}\end{minipage} - \hfill(\isa{findzero{\isachardot}pinduct})% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Remember that all we - are doing here is use some tricks to make a total function appear - as if it was partial. We can still write the term \isa{findzero\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isacharparenright}\ {\isadigit{0}}} and like any other term of type \isa{nat} it is equal - to some natural number, although we might not be able to find out - which one. The function is \emph{underdefined}. - - But it is defined enough to prove something interesting about it. We - can prove that if \isa{findzero\ f\ n} - terminates, it indeed returns a zero of \isa{f}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -\noindent We apply induction as usual, but using the partial induction - rule:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ f\ n\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}% -\begin{isamarkuptxt}% -\noindent This gives the following subgoals: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ }{\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}% -\end{isabelle} - - \noindent The hypothesis in our lemma was used to satisfy the first premise in - the induction rule. However, we also get \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} as a local assumption in the induction step. This - allows to unfold \isa{findzero\ f\ n} using the \isa{psimps} - rule, and the rest is trivial. Since the \isa{psimps} rules carry the - \isa{{\isacharbrackleft}simp{\isacharbrackright}} attribute by default, we just need a single step:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ simp\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Proofs about partial functions are often not harder than for total - functions. Fig.~\ref{findzero_isar} shows a slightly more - complicated proof written in Isar. It is verbose enough to show how - partiality comes into play: From the partial induction, we get an - additional domain condition hypothesis. Observe how this condition - is applied when calls to \isa{findzero} are unfolded.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ f\ n\ \isacommand{assume}\isamarkupfalse% -\ dom{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ IH{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ x{\isacharunderscore}range{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ dom\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ x{\isacharunderscore}range\ \isacommand{show}\isamarkupfalse% -\ False\ \isacommand{by}\isamarkupfalse% -\ auto\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ x{\isacharunderscore}range\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isacharequal}\ n\ {\isasymor}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ auto\isanewline -\ \ \isacommand{thus}\isamarkupfalse% -\ {\isachardoublequoteopen}f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ dom\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ IH\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{A proof about a partial function}\label{findzero_isar} -\end{figure} -% -\isamarkupsubsection{Partial termination proofs% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Now that we have proved some interesting properties about our - function, we should turn to the domain predicate and see if it is - actually true for some values. Otherwise we would have just proved - lemmas with \isa{False} as a premise. - - Essentially, we need some introduction rules for \isa{findzero{\isacharunderscore}dom}. The function package can prove such domain - introduction rules automatically. But since they are not used very - often (they are almost never needed if the function is total), this - functionality is disabled by default for efficiency reasons. So we have to go - back and ask for them explicitly by passing the \isa{{\isacharparenleft}domintros{\isacharparenright}} option to the function package: - -\vspace{1ex} -\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% -\cmd{where}\isanewline% -\ \ \ldots\\ - - \noindent Now the package has proved an introduction rule for \isa{findzero{\isacharunderscore}dom}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ findzero{\isachardot}domintros% -\begin{isamarkuptext}% -\begin{isabelle}% -{\isacharparenleft}{\isadigit{0}}\ {\isacharless}\ {\isacharquery}f\ {\isacharquery}n\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% -\end{isabelle} - - Domain introduction rules allow to show that a given value lies in the - domain of a function, if the arguments of all recursive calls - are in the domain as well. They allow to do a \qt{single step} in a - termination proof. Usually, you want to combine them with a suitable - induction principle. - - Since our function increases its argument at recursive calls, we - need an induction principle which works \qt{backwards}. We will use - \isa{inc{\isacharunderscore}induct}, which allows to do induction from a fixed number - \qt{downwards}: - - \begin{center}\isa{{\isasymlbrakk}{\isacharquery}i\ {\isasymle}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharquery}j{\isacharsemicolon}\ {\isasymAnd}i{\isachardot}\ {\isasymlbrakk}i\ {\isacharless}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharparenleft}Suc\ i{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ i{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}i}\hfill(\isa{inc{\isacharunderscore}induct})\end{center} - - Figure \ref{findzero_term} gives a detailed Isar proof of the fact - that \isa{findzero} terminates if there is a zero which is greater - or equal to \isa{n}. First we derive two useful rules which will - solve the base case and the step case of the induction. The - induction is then straightforward, except for the unusual induction - principle.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}termination{\isacharcolon}\isanewline -\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isasymge}\ n{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\ \isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ base{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ {\isacharparenleft}simp\ add{\isacharcolon}{\isacharbackquoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isacharbackquoteclose}{\isacharparenright}\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ step{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}i{\isachardot}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}\ \isanewline -\ \ \ \ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ simp\isanewline -\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isacharbackquoteopen}x\ {\isasymge}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ base{\isacharparenright}\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ i\ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{thus}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ step{\isacharparenright}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{Termination proof for \isa{findzero}}\label{findzero_term} -\end{figure} -% -\begin{isamarkuptext}% -Again, the proof given in Fig.~\ref{findzero_term} has a lot of - detail in order to explain the principles. Using more automation, we - can also have a short proof:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}termination{\isacharunderscore}short{\isacharcolon}\isanewline -\ \ \isakeyword{assumes}\ zero{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isachargreater}{\isacharequal}\ n{\isachardoublequoteclose}\ \isanewline -\ \ \isakeyword{assumes}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{using}\isamarkupfalse% -\ zero\isanewline -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\ {\isacharparenleft}auto\ intro{\isacharcolon}\ findzero{\isachardot}domintros{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent It is simple to combine the partial correctness result with the - termination lemma:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}total{\isacharunderscore}correctness{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}blast\ intro{\isacharcolon}\ findzero{\isacharunderscore}zero\ findzero{\isacharunderscore}termination{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Definition of the domain predicate% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Sometimes it is useful to know what the definition of the domain - predicate looks like. Actually, \isa{findzero{\isacharunderscore}dom} is just an - abbreviation: - - \begin{isabelle}% -findzero{\isacharunderscore}dom\ {\isasymequiv}\ accp\ findzero{\isacharunderscore}rel% -\end{isabelle} - - The domain predicate is the \emph{accessible part} of a relation \isa{findzero{\isacharunderscore}rel}, which was also created internally by the function - package. \isa{findzero{\isacharunderscore}rel} is just a normal - inductive predicate, so we can inspect its definition by - looking at the introduction rules \isa{findzero{\isacharunderscore}rel{\isachardot}intros}. - In our case there is just a single rule: - - \begin{isabelle}% -{\isacharquery}f\ {\isacharquery}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}rel\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% -\end{isabelle} - - The predicate \isa{findzero{\isacharunderscore}rel} - describes the \emph{recursion relation} of the function - definition. The recursion relation is a binary relation on - the arguments of the function that relates each argument to its - recursive calls. In general, there is one introduction rule for each - recursive call. - - The predicate \isa{findzero{\isacharunderscore}dom} is the accessible part of - that relation. An argument belongs to the accessible part, if it can - be reached in a finite number of steps (cf.~its definition in \isa{Wellfounded{\isachardot}thy}). - - Since the domain predicate is just an abbreviation, you can use - lemmas for \isa{accp} and \isa{findzero{\isacharunderscore}rel} directly. Some - lemmas which are occasionally useful are \isa{accpI}, \isa{accp{\isacharunderscore}downward}, and of course the introduction and elimination rules - for the recursion relation \isa{findzero{\isachardot}intros} and \isa{findzero{\isachardot}cases}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{A Useful Special Case: Tail recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The domain predicate is our trick that allows us to model partiality - in a world of total functions. The downside of this is that we have - to carry it around all the time. The termination proof above allowed - us to replace the abstract \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} by the more - concrete \isa{n\ {\isasymle}\ x\ {\isasymand}\ f\ x\ {\isacharequal}\ {\isadigit{0}}}, but the condition is still - there and can only be discharged for special cases. - In particular, the domain predicate guards the unfolding of our - function, since it is there as a condition in the \isa{psimp} - rules. - - Now there is an important special case: We can actually get rid - of the condition in the simplification rules, \emph{if the function - is tail-recursive}. The reason is that for all tail-recursive - equations there is a total function satisfying them, even if they - are non-terminating. - -% A function is tail recursive, if each call to the function is either -% equal -% -% So the outer form of the -% -%if it can be written in the following -% form: -% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} - - - The function package internally does the right construction and can - derive the unconditional simp rules, if we ask it to do so. Luckily, - our \isa{findzero} function is tail-recursive, so we can just go - back and add another option to the \cmd{function} command: - -\vspace{1ex} -\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharcomma}\ tailrec{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% -\cmd{where}\isanewline% -\ \ \ldots\\% - - - \noindent Now, we actually get unconditional simplification rules, even - though the function is partial:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ findzero{\isachardot}simps% -\begin{isamarkuptext}% -\begin{isabelle}% -findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% -\end{isabelle} - - \noindent Of course these would make the simplifier loop, so we better remove - them from the simpset:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\isamarkupfalse% -\ findzero{\isachardot}simps{\isacharbrackleft}simp\ del{\isacharbrackright}% -\begin{isamarkuptext}% -Getting rid of the domain conditions in the simplification rules is - not only useful because it simplifies proofs. It is also required in - order to use Isabelle's code generator to generate ML code - from a function definition. - Since the code generator only works with equations, it cannot be - used with \isa{psimp} rules. Thus, in order to generate code for - partial functions, they must be defined as a tail recursion. - Luckily, many functions have a relatively natural tail recursive - definition.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Nested recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Recursive calls which are nested in one another frequently cause - complications, since their termination proof can depend on a partial - correctness property of the function itself. - - As a small example, we define the \qt{nested zero} function:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ nz\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}nz\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}nz\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ nz\ {\isacharparenleft}nz\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -If we attempt to prove termination using the identity measure on - naturals, this fails:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ \isacommand{apply}\isamarkupfalse% -\ auto% -\begin{isamarkuptxt}% -We get stuck with the subgoal - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharless}\ Suc\ n% -\end{isabelle} - - Of course this statement is true, since we know that \isa{nz} is - the zero function. And in fact we have no problem proving this - property by induction.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isacommand{lemma}\isamarkupfalse% -\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}nz{\isachardot}pinduct{\isacharparenright}\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -We formulate this as a partial correctness lemma with the condition - \isa{nz{\isacharunderscore}dom\ n}. This allows us to prove it with the \isa{pinduct} rule before we have proved termination. With this lemma, - the termination proof works as expected:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ {\isacharparenleft}auto\ simp{\isacharcolon}\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -As a general strategy, one should prove the statements needed for - termination as a partial property first. Then they can be used to do - the termination proof. This also works for less trivial - examples. Figure \ref{f91} defines the 91-function, a well-known - challenge problem due to John McCarthy, and proves its termination.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -\isacommand{function}\isamarkupfalse% -\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n\ then\ n\ {\isacharminus}\ {\isadigit{1}}{\isadigit{0}}\ else\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate{\isacharcolon}\ \isanewline -\ \ \isakeyword{assumes}\ trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ n{\isachardoublequoteclose}\ \isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}n\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{using}\isamarkupfalse% -\ trm\ \isacommand{by}\isamarkupfalse% -\ induct\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{termination}\isamarkupfalse% -\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{let}\isamarkupfalse% -\ {\isacharquery}R\ {\isacharequal}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{1}}\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}wf\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isacharcolon}{\isacharcolon}\ nat\ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isachardoublequoteclose}\ % -\isamarkupcmt{Assumptions for both calls% -} -\isanewline -\isanewline -\ \ \isacommand{thus}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\ % -\isamarkupcmt{Inner call% -} -\isanewline -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ inner{\isacharunderscore}trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ % -\isamarkupcmt{Outer call% -} -\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ {\isacharbackquoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupfalse\isabellestyle{tt} -\end{minipage} -\vspace{6pt}\hrule -\caption{McCarthy's 91-function}\label{f91} -\end{figure} -% -\isamarkupsection{Higher-Order Recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Higher-order recursion occurs when recursive calls - are passed as arguments to higher-order combinators such as \isa{map}, \isa{filter} etc. - As an example, imagine a datatype of n-ary trees:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isacharprime}a\ tree\ {\isacharequal}\ \isanewline -\ \ Leaf\ {\isacharprime}a\ \isanewline -{\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ list{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We can define a function which swaps the left and right subtrees recursively, using the - list functions \isa{rev} and \isa{map}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ mirror\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Leaf\ n{\isacharparenright}\ {\isacharequal}\ Leaf\ n{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Branch\ l{\isacharparenright}\ {\isacharequal}\ Branch\ {\isacharparenleft}rev\ {\isacharparenleft}map\ mirror\ l{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Although the definition is accepted without problems, let us look at the termination proof:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -% -\begin{isamarkuptxt}% -As usual, we have to give a wellfounded relation, such that the - arguments of the recursive calls get smaller. But what exactly are - the arguments of the recursive calls when mirror is given as an - argument to \isa{map}? Isabelle gives us the - subgoals - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharquery}R\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}l\ x{\isachardot}\ x\ {\isasymin}\ set\ l\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ Branch\ l{\isacharparenright}\ {\isasymin}\ {\isacharquery}R% -\end{isabelle} - - So the system seems to know that \isa{map} only - applies the recursive call \isa{mirror} to elements - of \isa{l}, which is essential for the termination proof. - - This knowledge about \isa{map} is encoded in so-called congruence rules, - which are special theorems known to the \cmd{function} command. The - rule for \isa{map} is - - \begin{isabelle}% -{\isasymlbrakk}{\isacharquery}xs\ {\isacharequal}\ {\isacharquery}ys{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ set\ {\isacharquery}ys\ {\isasymLongrightarrow}\ {\isacharquery}f\ x\ {\isacharequal}\ {\isacharquery}g\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ map\ {\isacharquery}f\ {\isacharquery}xs\ {\isacharequal}\ map\ {\isacharquery}g\ {\isacharquery}ys% -\end{isabelle} - - You can read this in the following way: Two applications of \isa{map} are equal, if the list arguments are equal and the functions - coincide on the elements of the list. This means that for the value - \isa{map\ f\ l} we only have to know how \isa{f} behaves on - the elements of \isa{l}. - - Usually, one such congruence rule is - needed for each higher-order construct that is used when defining - new functions. In fact, even basic functions like \isa{If} and \isa{Let} are handled by this mechanism. The congruence - rule for \isa{If} states that the \isa{then} branch is only - relevant if the condition is true, and the \isa{else} branch only if it - is false: - - \begin{isabelle}% -{\isasymlbrakk}{\isacharquery}b\ {\isacharequal}\ {\isacharquery}c{\isacharsemicolon}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}u{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}v{\isasymrbrakk}\isanewline -{\isasymLongrightarrow}\ {\isacharparenleft}if\ {\isacharquery}b\ then\ {\isacharquery}x\ else\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}c\ then\ {\isacharquery}u\ else\ {\isacharquery}v{\isacharparenright}% -\end{isabelle} - - Congruence rules can be added to the - function package by giving them the \isa{fundef{\isacharunderscore}cong} attribute. - - The constructs that are predefined in Isabelle, usually - come with the respective congruence rules. - But if you define your own higher-order functions, you may have to - state and prove the required congruence rules yourself, if you want to use your - functions in recursive definitions.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Congruence Rules and Evaluation Order% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Higher order logic differs from functional programming languages in - that it has no built-in notion of evaluation order. A program is - just a set of equations, and it is not specified how they must be - evaluated. - - However for the purpose of function definition, we must talk about - evaluation order implicitly, when we reason about termination. - Congruence rules express that a certain evaluation order is - consistent with the logical definition. - - Consider the following function.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{0}}\ {\isasymor}\ f\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -For this definition, the termination proof fails. The default configuration - specifies no congruence rule for disjunction. We have to add a - congruence rule that specifies left-to-right evaluation order: - - \vspace{1ex} - \noindent \isa{{\isasymlbrakk}{\isacharquery}P\ {\isacharequal}\ {\isacharquery}P{\isacharprime}{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}P{\isacharprime}\ {\isasymLongrightarrow}\ {\isacharquery}Q\ {\isacharequal}\ {\isacharquery}Q{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharquery}P\ {\isasymor}\ {\isacharquery}Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}P{\isacharprime}\ {\isasymor}\ {\isacharquery}Q{\isacharprime}{\isacharparenright}}\hfill(\isa{disj{\isacharunderscore}cong}) - \vspace{1ex} - - Now the definition works without problems. Note how the termination - proof depends on the extra condition that we get from the congruence - rule. - - However, as evaluation is not a hard-wired concept, we - could just turn everything around by declaring a different - congruence rule. Then we can make the reverse definition:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ disj{\isacharunderscore}cong{\isadigit{2}}{\isacharbrackleft}fundef{\isacharunderscore}cong{\isacharbrackright}{\isacharcolon}\ \isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isasymnot}\ Q{\isacharprime}\ {\isasymLongrightarrow}\ P\ {\isacharequal}\ P{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}Q\ {\isacharequal}\ Q{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}P{\isacharprime}\ {\isasymor}\ Q{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{fun}\isamarkupfalse% -\ f{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}f{\isacharprime}\ n\ {\isacharequal}\ {\isacharparenleft}f{\isacharprime}\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymor}\ n\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent These examples show that, in general, there is no \qt{best} set of - congruence rules. - - However, such tweaking should rarely be necessary in - practice, as most of the time, the default set of congruence rules - works well.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/Thy/document/session.tex --- a/doc-src/IsarAdvanced/Functions/Thy/document/session.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ -\input{Functions.tex} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/conclusion.tex --- a/doc-src/IsarAdvanced/Functions/conclusion.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,7 +0,0 @@ -\section{Conclusion} - -\fixme{} - - - - diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/functions.tex --- a/doc-src/IsarAdvanced/Functions/functions.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,91 +0,0 @@ - -\documentclass[a4paper,fleqn]{article} - -\usepackage{latexsym,graphicx} -\usepackage[refpage]{nomencl} -\usepackage{../../iman,../../extra,../../isar,../../proof} -\usepackage{../../isabelle,../../isabellesym} -\usepackage{style} -\usepackage{mathpartir} -\usepackage{amsthm} -\usepackage{../../pdfsetup} - -\newcommand{\cmd}[1]{\isacommand{#1}} - -\newcommand{\isasymINFIX}{\cmd{infix}} -\newcommand{\isasymLOCALE}{\cmd{locale}} -\newcommand{\isasymINCLUDES}{\cmd{includes}} -\newcommand{\isasymDATATYPE}{\cmd{datatype}} -\newcommand{\isasymAXCLASS}{\cmd{axclass}} -\newcommand{\isasymDEFINES}{\cmd{defines}} -\newcommand{\isasymNOTES}{\cmd{notes}} -\newcommand{\isasymCLASS}{\cmd{class}} -\newcommand{\isasymINSTANCE}{\cmd{instance}} -\newcommand{\isasymLEMMA}{\cmd{lemma}} -\newcommand{\isasymPROOF}{\cmd{proof}} -\newcommand{\isasymQED}{\cmd{qed}} -\newcommand{\isasymFIX}{\cmd{fix}} -\newcommand{\isasymASSUME}{\cmd{assume}} -\newcommand{\isasymSHOW}{\cmd{show}} -\newcommand{\isasymNOTE}{\cmd{note}} -\newcommand{\isasymCODEGEN}{\cmd{code\_gen}} -\newcommand{\isasymPRINTCODETHMS}{\cmd{print\_codethms}} -\newcommand{\isasymFUN}{\cmd{fun}} -\newcommand{\isasymFUNCTION}{\cmd{function}} -\newcommand{\isasymPRIMREC}{\cmd{primrec}} -\newcommand{\isasymRECDEF}{\cmd{recdef}} - -\newcommand{\qt}[1]{``#1''} -\newcommand{\qtt}[1]{"{}{#1}"{}} -\newcommand{\qn}[1]{\emph{#1}} -\newcommand{\strong}[1]{{\bfseries #1}} -\newcommand{\fixme}[1][!]{\strong{FIXME: #1}} - -\newtheorem{exercise}{Exercise}{\bf}{\itshape} -%\newtheorem*{thmstar}{Theorem}{\bf}{\itshape} - -\hyphenation{Isabelle} -\hyphenation{Isar} - -\isadroptag{theory} -\title{Defining Recursive Functions in Isabelle/HOL} -\author{Alexander Krauss} - -\isabellestyle{tt} -\renewcommand{\isastyletxt}{\isastyletext}% use same formatting for txt and text - -\begin{document} - -\date{\ \\} -\maketitle - -\begin{abstract} - This tutorial describes the use of the new \emph{function} package, - which provides general recursive function definitions for Isabelle/HOL. - We start with very simple examples and then gradually move on to more - advanced topics such as manual termination proofs, nested recursion, - partiality, tail recursion and congruence rules. -\end{abstract} - -%\thispagestyle{empty}\clearpage - -%\pagenumbering{roman} -%\clearfirst - -\input{intro.tex} -\input{Thy/document/Functions.tex} -%\input{conclusion.tex} - -\begingroup -%\tocentry{\bibname} -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{../../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/intro.tex --- a/doc-src/IsarAdvanced/Functions/intro.tex Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,55 +0,0 @@ -\section{Introduction} - -Starting from Isabelle 2007, new facilities for recursive -function definitions~\cite{krauss2006} are available. They provide -better support for general recursive definitions than previous -packages. But despite all tool support, function definitions can -sometimes be a difficult thing. - -This tutorial is an example-guided introduction to the practical use -of the package and related tools. It should help you get started with -defining functions quickly. For the more difficult definitions we will -discuss what problems can arise, and how they can be solved. - -We assume that you have mastered the fundamentals of Isabelle/HOL -and are able to write basic specifications and proofs. To start out -with Isabelle in general, consult the Isabelle/HOL tutorial -\cite{isa-tutorial}. - - - -\paragraph{Structure of this tutorial.} -Section 2 introduces the syntax and basic operation of the \cmd{fun} -command, which provides full automation with reasonable default -behavior. The impatient reader can stop after that -section, and consult the remaining sections only when needed. -Section 3 introduces the more verbose \cmd{function} command which -gives fine-grained control. This form should be used -whenever the short form fails. -After that we discuss more specialized issues: -termination, mutual, nested and higher-order recursion, partiality, pattern matching -and others. - - -\paragraph{Some background.} -Following the LCF tradition, the package is realized as a definitional -extension: Recursive definitions are internally transformed into a -non-recursive form, such that the function can be defined using -standard definition facilities. Then the recursive specification is -derived from the primitive definition. This is a complex task, but it -is fully automated and mostly transparent to the user. Definitional -extensions are valuable because they are conservative by construction: -The \qt{new} concept of general wellfounded recursion is completely reduced -to existing principles. - - - - -The new \cmd{function} command, and its short form \cmd{fun} have mostly -replaced the traditional \cmd{recdef} command \cite{slind-tfl}. They solve -a few of technical issues around \cmd{recdef}, and allow definitions -which were not previously possible. - - - - diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/isabelle_isar.eps --- a/doc-src/IsarAdvanced/Functions/isabelle_isar.eps Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6488 +0,0 @@ -%!PS-Adobe-2.0 EPSF-1.2 -%%Title: isabelle_any -%%Creator: FreeHand 5.5 -%%CreationDate: 24.09.1998 21:04 Uhr -%%BoundingBox: 0 0 202 178 -%%FHPathName:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any -%ALDOriginalFile:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any -%ALDBoundingBox: -153 -386 442 456 -%%FHPageNum:1 -%%DocumentSuppliedResources: procset Altsys_header 4 0 -%%ColorUsage: Color -%%DocumentProcessColors: Cyan Magenta Yellow Black -%%DocumentNeededResources: font Symbol -%%+ font ZapfHumanist601BT-Bold -%%DocumentFonts: Symbol -%%+ ZapfHumanist601BT-Bold -%%DocumentNeededFonts: Symbol -%%+ ZapfHumanist601BT-Bold -%%EndComments -%!PS-AdobeFont-1.0: ZapfHumanist601BT-Bold 003.001 -%%CreationDate: Mon Jun 22 16:09:28 1992 -%%VMusage: 35200 38400 -% Bitstream Type 1 Font Program -% Copyright 1990-1992 as an unpublished work by Bitstream Inc., Cambridge, MA. -% All rights reserved. -% Confidential and proprietary to Bitstream Inc. -% U.S. GOVERNMENT RESTRICTED RIGHTS -% This software typeface product is provided with RESTRICTED RIGHTS. Use, -% duplication or disclosure by the Government is subject to restrictions -% as set forth in the license agreement and in FAR 52.227-19 (c) (2) (May, 1987), -% when applicable, or the applicable provisions of the DOD FAR supplement -% 252.227-7013 subdivision (a) (15) (April, 1988) or subdivision (a) (17) -% (April, 1988). Contractor/manufacturer is Bitstream Inc., -% 215 First Street, Cambridge, MA 02142. -% Bitstream is a registered trademark of Bitstream Inc. -11 dict begin -/FontInfo 9 dict dup begin - /version (003.001) readonly def - /Notice (Copyright 1990-1992 as an unpublished work by Bitstream Inc. All rights reserved. Confidential.) readonly def - /FullName (Zapf Humanist 601 Bold) readonly def - /FamilyName (Zapf Humanist 601) readonly def - /Weight (Bold) readonly def - /ItalicAngle 0 def - /isFixedPitch false def - /UnderlinePosition -136 def - /UnderlineThickness 85 def -end readonly def -/FontName /ZapfHumanist601BT-Bold def -/PaintType 0 def -/FontType 1 def -/FontMatrix [0.001 0 0 0.001 0 0] readonly def -/Encoding StandardEncoding def -/FontBBox {-167 -275 1170 962} readonly def -/UniqueID 15530396 def -currentdict end -currentfile eexec -a2951840838a4133839ca9d22e2b99f2b61c767cd675080aacfcb24e19cd -1336739bb64994c56737090b4cec92c9945ff0745ef7ffc61bb0a9a3b849 -e7e98740e56c0b5af787559cc6956ab31e33cf8553d55c0b0e818ef5ec6b -f48162eac42e7380ca921dae1c82b38fd6bcf2001abb5d001a56157094cf -e27d8f4eac9693e88372d20358b47e0c3876558ebf757a1fbc5c1cddf62b -3c57bf727ef1c4879422c142a084d1c7462ac293e097fabe3a3ecfcd8271 -f259833bac7912707218ec9a3063bf7385e02d8c1058ac06df00b33b8c01 -8768b278010eb4dd58c7ba59321899741cb7215d8a55bee8d3398c887f02 -e1f4869387f89141de693fcb429c7884c22dcdeddcaa62b7f5060249dfab -cfc351201f7d188b6ed68a228abda4d33b3d269ac09cde172bc045e67449 -c0f25d224efbe8c9f9d2968a01edbfb039123c365ed0db58ad38aabe015b -8881191dd23092f6d53d5c1cd68ebd038e098d32cb24b433b7d5d89c28ee -05ea0b6070bb785a2974b5a160ee4cf8b6d8c73445d36720af0530441cd9 -61bc0c367f1af1ec1c5ab7255ddda153c1868aba58cd5b44835535d85326 -5d7fed5ff7118adb5d5b76cc3b72e5ff27e21eb857261b3afb7688fca12d -1663b0d8bdc1dd47a84b65b47d3e76d3b8fa8b319f17e1bb22b45a7482fd -f9ad1b6129e09ae47f15cd2447484cd2d64f59ab0f2f876c81e7d87ccdf9 -005aa8fc093d02db51a075d571e925f2d309a1c535a1e59d34215c6cd33e -3c38997b4956461f376399901a8d0943dca6a335baac93fc8482c0659f04 -329c6f040b35828ea6dd1bd1858f2a9be4ef77731b5b75a1c536c6bc9479 -0821e5d88f6e2981835dbfd65ec254ebcf2cf49c917d121cd3bbb476a12b -69c15f17d9c17bb15ad1e7d31d2afcf58c8f0ad526d68615a0f1ac3b1d1c -d3beafeea3cf56c8f8a66367c70df9159f0b1b3157ccfd010045c4718e0e -625c0891e85790c9b97b85517c74c9d55eaca31a01cddc64898bf0eeadf3 -53391a185e507dcb0a6f52661a56357ac818dfc740a540aadf02f4e7a79d -8008a77cd30abde337025b01217d6a68c306abe145b7260c3478fa5f366f -b2d37259ead8a8ec2db2f09ae0eb3a682d27b0d73a60935f80254c24426a -003a87a29a4370cbf1b2ef1e19ad8466ec725fd5b463d06011a5e0da2258 -ff6c1483c4532bc21f2ed9b99b929b2800ddefc1a98d12ba085adc210bac -e0274b69e24d16af015a51ca73edf779a7534a887aa780337ad966839881 -edc22ba72038aa1a96a6deba24ad676795da711b92f8cf5f54cb4322ec04 -40ef9e15b11e3005f3ff69376ecb29bb66e8fc1b685f2b05fb162fcb35aa -d7eb2a8ec39b97ab1ff05ef02f8dbbc12085a5cd252cc4010fab7f63dfd5 -7fa1be86f724d37db5faef17ae8e8e537444e8e9db917b270344183473af -7f47d5703a160d8ef1e772438620d3336b2fbcf6433612e4b5e64fae0329 -8a3c4010c17d93f13ba66d549c69dd58c7d26ddc90285fed831918563a16 -2a7ac2511e2f18c9eb3df905a9dcba65a31cc1c39f07458abb11b4c60346 -aea19070e126982f1dde336e79be0ecd69a8afbe2493d064d4d2ff38788b -b3038125961302db9761403c3b8019ec641e107425002205a77ae2ae0f4f -7550d623dd03f0ec0199f42a9a8b89514e2e21baca9b3c8c96ca48cbf9dc -ee6d6241d713e014b49e83ad85e62a6b2f70b58e4cc72d41ea6fcbdd3b5c -890c8af0d24200658773b1628c6cc9aaaabb08865ee4c7ff2c513ad7aa23 -155a321213fa94731683a3e282a0e60aa3c87aade3da231465bdd9097f2c -89a1af8e5b9649143f2d9482546501ea97e8bea2f5d5eea97d4f19bb6835 -3138d3babb2461c08d537491aaede1f23d734c6f099eb5bef6e2ffaaf138 -e5ab71b8b41599091037e440127a4eaedf208c20c8a2fc62eadab191d1ab -4d5531f826aa6b9fff2797a7f54673e0a3fae09a93a0dfafb8b11d60dc69 -5acf9b7e1a47c31d0b5a0b85b7b50cddff5ac831651d9c7469c2139c7a89 -7d2f868f36c65156921803eccfdbdd1618595ab6d2a9230ef523a1b5ee51 -f2a0d200fc0e94aff7f546593ff2a3eb865d129895af01b8ab6e4616fe20 -9123b6e2b7e0817adc3cdb78ae8b0b1d75f2986ebd8fb24c4de92ac9e8c3 -6afa520636bcad2e6a03d11c268d97fa578561f6e7523e042c4cc73a0eac -7a841907450e83d8e7a8de4db5085f6e8b25dc85b59e018380f4b9523a7f -02cbeec989f0221b7681ec427519062b429dcd8fc2e4f81173519f88e2e4 -3798b90a84060920f6ae789afd6a9182e7fec87cd2d4235d37a65c3f3bcc -c742c89cbe5f3e2ba6c4f40ebba162e12569d20684cc167685285a087e7a -0a995fe1939bf25c09553512ba2cf77ef21d2ef8da1c73ba6e5826f4099e -27d8bc7b3545fc592b75228e70127c15a382774357457cd4586d80dc0bd6 -065aee32acfd5c0523303cece85a3dbf46853b917618c0330146f527c15b -dbb9f6526964368b2b8593eed1551dad75659565d54c6a0a52da7a8e366f -dd009ef853491c0fb995e19933cba1dbdc8902721c3ea6017ffdd5851cb8 -3c8bada46075ac121afe13a70e87529e40693157adcc999ed4657e017adf -f7dbac4bc0d204f204c6f47b769aaf714f9ec1d25226f24d0a1b53e28ac5 -374ab99755852c1431b2486df5fd637e2005a25303345a1c95a15a1189ba -f6f6883de1ad46d48427b137c2003d210ab2b2f5680f2633939f289d7553 -eb943adf8127f1c3ee7d6453b5566393700ad74ab86eb9a89f8b0380af55 -6b62f51b7dbd0c5dcc9a9fb658944d7ad5845d58dedc2d38200d0ef7cb0f -76041dc104ef3ab89c1dc2f6a75635d48051c8a7dd9f5e60253a53957ec8 -9d1266566d7ed20d79dfc2807b397d7cf056bdaccdb72528a88aa4987682 -c909b2fe1e35a71c2f29e89a2bf32173967e79610367ce4574ba6a1cc031 -cfb176fc0313f74f91a866ef9954b95b29caf917a6b919586f54d23cb7ce -23305886ae7760ebd6263df0d3c511ac7afc361df78bc2621f66d3268b99 -078fa59124f0eb9476496c938eb4584e87455dc6f2faa999e938460b31c6 -28021c652acfa12d4556aa4302bbcd043e60389480b796c3fc0b2e51b81e -c2afa4a34335318a1c5a842dcaa120df414acba2e79ab5cc61c99e98108c -5cb907a96b30d731131782f9df79aabfc16a20ace8852d047497982e11c8 -26321addf679de8a96a2d18743f6f2c3b2bc397370b10ad273fcfb76b48b -9dad27cf89ca2c370149cd48ab956e9bbce01cbf8e1f0c661b99cf19b36e -87b6165dd85ae3f3674525e17d85971093e110520d17f2d6253611e35ec9 -e17000e48e2926454c1e8eb4098e0115b49536f2c33505eb5e574f7a414b -e176398c5ddf6d846ea5ddf2a5e94c0422e0115c57a8c9e56bf8ba962c82 -698c96bd6138baaca7347e44352cc62df4eeba364954ad921a5a43280980 -264df4a7fb29d005423179f7bd1d98b4280d62ce23c927551f1ffc2b8f17 -0a9c23656c0c91b640cdcfdbd88089ffb28d3ac68bad25dbbed82c083696 -1f9f86a6183cc1061ffdb32279796569d05b31c946955402d0be1fb9f2bf -304d1ad8e1e357be49e6e2ee67f7f1e7bc699d46a5f7853fe659ba2e1930 -0d3e3ea658b9862701dcab08fdd23bf1d751777f25efbe9e02d12b5612b3 -c3fc2275190346b94ec4024e4ade08e54d75c0b4c48f4956b9182e8ce997 -74b54da4a9318c099d89f1ce3b6803a95f48b9fb8b845372be25e54478e8 -49e4707ea03a36e134efa661e4e6250d89649ae074cfd9d6b9e2071c1099 -3b8a5a5ebc3e1cb228c97565aef7f254e3f90af8a3dd281c83792755719d -c6a5b3bab4aa6be5afe9624050eee8dfb13b018f4088c932cd48ace38dfe -b1b4218dba8f7fada6628076acf1b54db0c95d4fb12232f1fa9d1ba848f9 -fe80c65b75d6946c00fe78839142c5302707d9269b24565dbcc551aca060 -b4d0f99b961dd3cc795a982063ac42e9fc81fc98add42744a9f92e74b00d -637ee4606ea2099b6c763493c3159f8e52a90dafca682010c0e92bc9038a -10abb066c75c8d97f7ad6fb1a37136e52cf2093c4fa485fe12adad10e4d0 -83b78b023628ddc5326cbf8392516027a9f3de4945f93488e4a1997efd2a -22c2c136dbac1bdb829e082beac48cdd06dcb17bacf09451c7b636bd49a8 -fc60cb1d2292250bea78e1dd276725ab4c526b66ddabf4e1b2bf0a2571df -2490df70735f5da321fac74fe4fab54444e53dace9832cff38a70c58104a -4f0c0545dcf7a1a9ecb54e0e32d6d8195d30b2c98a567741fcf325a4ddeb -244a1a35676e246ddc835fac13b569f35f22ec93191eca3efbe17ff9a950 -d08f863710b2bbecec969068c498eb2338b68f3fc3f5447449fe4de2b065 -e068ecd9255d369b2bb6c4b0b7423774bed69294758aca2bdb1a8d5bf618 -d3fa09462f7838e8a79b7a53bebe6dacb0a1561eaa074bc6f069f6a06fb2 -c4a5cb13e2172bce9be659a82665da5cded73da84322bb16aa6e19ac1958 -7515cb5d2b65e65e803f76700ce5efd3df7fe4ed431fae0e8e286b1d5056 -a0d18df926b2be7a93c503ab903abd4788680a6201fdc299f2cb5d6a9b6e -2048109c8d1fb633a54128938594b2cce86a7e0185e7d59e6536584039ec -9e30ff7be6ddba9fdba82de7415fdc47de84d97afb1aa3ba495bd91dee9d -f3b21ee1164987dd8510925087cd0919f1085cba6e4dd3c7384d26667f94 -ad7f736a60d8bd76dfaa4b53c23536fc309ff2317c48ee0107ff2ca4d1b3 -f78c5a27b901c931128bdb636094ef0cd543a5b62b7dbe10ed83aed5780c -a16067a4a7e8b7c5bf8a8e822149bc1b5bcdabe13a7f6aa6eaeff24a42f4 -a58a2b70f545103540169137fda9abb589f734b6776cb50402d6123ce802 -10dce05e3697a98c9411cf60a02c278c91e03d540b936cd00c668960e357 -1aeaf4d94cfb496b259ec0d8fdba9199fb46634ff177bc8d310ea1314eef -d46c927a981c58e88743ed4e07d80fe841edee812e3053412bf2e710146c -b25dec8ea70c38bb1f6e4db3c2e8ba521963c1584eeb60ea1e9555058f13 -e98307c13cbd15c26b611f543149b1ddf88dd6296ae703f58daeb67f1b03 -ab5b24c72d5770cb9d8ed242c4faaad1dd940ada00e98ff3a61799d13355 -aba916910aa9a6e5ee8af438d0ba8235346fcd139b9d2cb7db7bd3f298a3 -94ff0aff3b9042f32a004e042c346a5ea35917f229848a9c5a32909b0090 -4aa715640277a6ada99f8b2927fda22899ff1347f42bac73e2bd4bbf3945 -55fd7dd30d5c6dadf5c259fdb2455d8c607de1c5da588e20159f18e4e8da -b714e532d888a0386c52c2b85964251af003ac9d10c0c8b9b3465e1dde48 -2e74a29e17a7cf6c9a43b5af1646f0b8508f98e9a1279ec3908073d89dcb -aa093e8dd1004c1ecccce0803095b0069d4be7a1eb14b02efc37d137dfe3 -f0190bc9628069abc257f45d0e050e60c7f5281277937dd986fcd5b94a2b -845a1a75addd74a142800f18ef408c08a2c2ad16a93298f148c0ae7d2990 -ded147f92f053809a60d4f992a227167aad5b5eb2bbe8a4a77dc77a08b72 -6acb34422e2532eec7e274e4a42d71ee584f5f3df5a3a5b379974ede73ab -5f1b0a2dbfcc8cfac4747ca26eb6030dc2f85a104b754432977850c093b9 -97ed90af711b544ff682e7b1eac82b2b7b44014b09c17ecf084c994a095d -9eeef5391305caf093b62ac9916f982a436d521fcf4d75c5b8e4d92266fd -e99a58aa39d7693ecd1796b8851761d64bbca39a6d5a0b4533ae47123327 -f98d0ad0e8b36625cc3647b55459552906d8a1d5766845ffac101980efcf -79657e365510be5db557cefef21193ca3cf3dad175ee2e7ae91d174fdc06 -2ff5c51ffe2f021122e96df042019d3a1883e662537ec1b69c11fbb6e750 -0132eabf5803c816153ecbff60ca3b3b39708c26cb1751afb2e65d8e5f4a -c4397a88fb1f112672fcdd24e4ba545c5b2a7968c17b62f8e2530a8acbff -cfca82c64b7abcab84e2c4a0a7ced67b15669301fe9ff2c756e70ff7ce33 -497be6acc4ac5617e1f043bd8a87416299a39bf17fc31c02d72d75fdc2a1 -e60669fa4d5e4a49d9afea2f53f4626680e9c0dfca223529efa415c4343a -b6067aa800c484457ea050eaaa5d3fafeedd0eec72f327e02c6b3912b5a8 -c404de4839c9c4a99da42681cde43316606a34c7d2f02269de1aab776857 -e668f35946af4d618d36d444bdc02b1f63ea25b6260b4fb606ac8575b5c9 -782a5de4037350d5753b1537537ccb008c454eeb264e6cd4727c999e240e -0ac89e95a896b67d54910a3531345f64198ad394b5ceb52881f1dd9e6beb -95862dc188d45b3e46aacb5fe40097947dab9bc3c1ee46bfc9b1b3ed6167 -efd0d65ceb043d7b24c1456676e4baa47b1209a315f199bb3a91f4374cd9 -cc0b40d3f09f19f8dd8a46915eee55eeeeb3c7b8f437106ee491ef0f4ff9 -2c5c6f779e0fbe7bd5293964bb645ca362b106abeb773571d9ae83b786a3 -d5a4ea3ea970daadc46cc5e6037f76fd20e0fffc47cf4e7af9522b91f96b -3297720fd45d9bc2200622ad2ca9445556c8a8202b1991bc63da360d021d -55be2528e043f803e08da99b91ab9cfc5e65b2655d78206b4aecd445a7b0 -1caa0d06b0a55e8f04b70b60b04a860c8e1ab439f4910051e3f7441b47c7 -8aa3ab8519f181a9e833f3242fa58d02ed76bf0031f71f9def8484ecced8 -b6e41aca56176b6b32a2443d12492c8a0f5ba8a3e227219dfd1dd23fcb48 -fcfd255dbbf3e198874e607399db8d8498e719f00e9ed8bdd96c88817606 -357a0063c23854e64ad4e59ddd5060845b2c4cddd00c40081458f8ee02c7 -303c11747bd104440046bf2d09794fca2c4beb23ed1b66d9ccb9a4dd57ad -a24943461ecc00704c916bdc621bfddb17913dfb0f3513b65f3ab015786a -caa51ee9546bc8ddf87e2e104137e35ddf8f8d23724e9a53824169bc7cfa -99562656e6f1c888d4dbff0b269c5d1e733e5f212d91297610201eb43249 -35e336dd0052738db2d64f3e89429903bb5c1810009cf766e9a06223dd2f -219b706394a121dc029af55c6ada9052af59682ef7c51e121cf16f0319ac -0aa9512ef900c548d673fe361da19052808797e958209072e145d46ec8cc -a89fafd76630eff30ae979973bdf0f8c9e469d8edd3b1c93731c72f976b7 -d81142bc15c376403f967affaa5f482efd57c6f91970729d16db851f0ed3 -ea7d82f409307b5b436886c1beda94a1fa3ab1b60686f6574c844fb2c0b3 -a07174dc4f27b4fed2f8bd4d5436be4b343e5efdf0400d235bd910255341 -a20770804a26f8437e9bce6da8e9f8258a343c7aee291f1510be306ae67a -ab1d7696453530c02fd153bbe49dbf62baad6146029cbd1656cdb76c952c -b93edfee76fe33832930be59636bb947e8ad285f20f663cccf484fba97d6 -7446c7b6c6f5857428bb1737d9ae801df75d9cb4d7bd59ef7a4cbadde928 -38f15d232005585d2e40483d2d3e08cc8f398bb43afedb84343c3ba3835d -0ba82a86dce859cf655f85e63e41365e0dbefcf511b9a27a2b6e66b2ad3a -c657902842287a317e46ceaa93b5088f09d53a65815b44538af90ad3b06b -4e5e2dc509f02e30a01e05201c67d4d39582bbe64e20b669f5fd787909a3 -30fc50a95b31426bbb57a4fbf9feacdc31f98bcf50da7e50c2bfc169c6fd -ccf213cdb878653bcea372968ea6e31fd30dd55434cc91c0af22179ce669 -a05493f195e12432c6173ae2ac3c94fb83f38210014a9f969ea2b44e99f5 -e5a7317e848d429ad62167a4fc5001149676c0c28cdf59b8d1c5a582f516 -3eee855312777fee6dacbf993f5c058f355dbde6552dc960d336eee445dd -11d53fd21b745d1e5ec317efbbef25e070d0a36797a6081c356ac2328e64 -e5c55fbc81dc75d9c1575548ece74b8307eef485aa8e28859a2e0435c831 -23a600efb323c362fe9f02407a5411c41a69566cd50add324b63ab939980 -b9d7a929ae4887163cfa7acbfc9fabaab8987a1f6906b9881491cd055b94 -485c968479dbb05b34ed0cd6844729a692978c6928c3392e33e8324ded88 -814cacdac8128e1425c0091a13558100d7cdbed5992795d94d39c32f32dc -621ab6f3b75187a66741f61d6a9c91d791b1cfc3d0e94d4a76302e0c3f2e -cbdc51f14f3251aa5c8bb989f0e13ee500b7b7f2f1e52ca970ad8a7b4b99 -57e93126254297380d67179deb8ff1e99d5cdf7a35c5bb9fa7b402e63234 -78640344e1f10c378ad23c5cd1aa18e1e0b308db70d3a624a455f8e291a2 -ee102ad10776208c2d546cb76d89ca8103a8b95f8acc2d2bdc9791324915 -6c9e05429091071f0c5b76d82c8d1c8a69d840fd460922cd2090624bc218 -0c9391005926a25042a55e322060807363462e1cdeab309033124ba3a884 -1db13f39edae04ec52cde9dbde64ddda1ad805141d4230ec76bd81fd98d7 -0d90fa1aaa26ea551bf687ddd6cdcf3de5a446b266c68434f07d9c0b382d -5816c4e22f22cc03ff78064c6dffb12315c6bcbbf5dc510f5aaabf23471a -234efceeb4aa2f9af9ea787c014c5587ef162fc5b35e8f4c23b168c6e247 -41d33dcc11d2a56d3ba9d8eed6e79aebf9f0faf1a3aeb89d792d69041f0b -b8fadfc0aa090effc6ae5e2f13cdbf54b5bed69b039eef2627769613b6f1 -aefe9b66747fe8feaf7455796740f411a770d4a1764f0483719584880f45 -430e38d3af184145892a08b2add234a3f3ee4ccfc9f6995c02392adafccd -722f366d748cfe9373fbf5f878ed47e9d221fd156bb28369df9e7d2b79da -76120d135ebaf36cff93beb7e313c2b2de7477176fc19609a1b906c995cd -defef08899265b6b8aefb44da1aadefd1c523dce5ca1b84c0c652b3009fd -057789892d4d31764f181754b2e0a62c465587585509989a219711a5e4e2 -5b3b340ca8fdd3f04fef204b1b722b2f6c2ccb00c3cf1a94ba9bdfbfeda9 -e2a062c6f1ced3b8aae5dae32ade1fca1001f98d0ad0e8b36625cc3647b5 -5459552906d8a788eb8bc734ccb65fe9582c71df94fd95d22c5323de235c -28220fb9a2ccb37362174d8cd5922c9e5a87b51d0668555100a33e33750e -f1f795cbed962494a994be7ce8cf71fc58ff4204551b1615ed27cf088171 -fd000b72462b67935961e7c6c3a05bfd67b9ba094ea2c16fdf486da912e1 -e97bfd1c17934535e551cede20c001b5d2adb2be4cbad7d6ba0bdeae4b1a -a739f90293e67ecbdeea4d35825e092697fb05b215083e3f3d6be260790e -2a175fd44eb1c4c16759504827a6eb58a838c4d65fec6eef108495577019 -15740cac164111892e8d1cc447cd208e243a89ab847d8ebf4fb98bff49e7 -a3453facf3b0e8cb67590f390173ddba68324531d2e426aed152e12301d7 -538c1f3c0048a9cc00c009a1a9138460082123209c1e007266fbf236eb72 -21f87d4ca38a0b699e84ca230ffb5095f90a6528bf2a9118f95ac9ab8d2d -ed9eed9b8b27be894b717469758c8d94fa89acc64f530f432d0e5f16c922 -36d6a63410f099c9e909450fd731d698ef658d8ffc1de14817b850814f68 -1a4a9be5cc7a71c381974c249f0b209bfdc2e97f9540c96f57bb4d283622 -00969b82011315289e6a025b137030a0af3b4b42b00fed7cec49df43c59e -3b2495a036dd1b17a8e6adae63bfbbd48807c44b5bbf71813355e1b0e58e -22b6fb88005fc55565be49c17244901b73ef02fc4eb7669be5af22d89c0d -dff0fc6821d810d13e5821d48d4a71d7e463d5b60bc279d0dcf5f8da3a95 -905b56d6f2be95e6d4243b1048e3b662e62401ffaa3bc3f5f90b0854b8a3 -8c38039f61fcb359b06bbb7d59e3b39a295dccd6db9a8b83a6f64ef8dc94 -a77123dd164cfd1c46f1ee51aa19c3d6e7db92a298d10159f2b5eff2caf9 -dc93a6d267fb65bd900d6adf0c6be598050b6d3a9b3a322ab3c9e880d774 -1f58016ff97e5f606b5dbd72ba99252c669209bb556dd5be84fdd7c1ce92 -8a3b3d3aab8d37e6b740227563bb4d60f6bb04052356e1a48d2079feca44 -7ea17fd06f208426d045dee660d1d6460455f8d20dbc5ae64550bbdf60d7 -27d96cc9afef842a8c8c78ea2257e6c6d0d207c80cfe399e8874c693274e -d2c2022d303ca50a70624b07434fb85040a76a823f446c7454dab4f9c05f -10274eb5ba164aa3649d1bc90694316ba5cb3e7df4442e777124cff7ebef -53df2320a0c441ab61666493cb43da46d5711c21699de85bc74359444da2 -e3e397d4c16234f81531505b621aa242a6698886f82b447104b1f1062f60 -b5c87cea9151bb3c627bfa4532b06fd147c556ed8d61ae30a8719dfb8705 -f8a6c74368381403640cc57026d3790c49e2bbd1c0e48285ec6ba44de678 -e3a1394d659c412f09644b83ee1a333a1f51ad8deb4e6d77b3b226ac2c4f -fe653411a7976ae7c4a3cb7df309788da6b483f8a7bab4a6990db74362f5 -bc41d545a320389b2599fd726e426ed9fa2916ece67b058f6a269544e517 -128bda38d117f402409d0d8f8c88ed509aa2ba882e0c579b45af4be80770 -22d7269684eaf0f9afc3054316da6611e3fd260d67fb6fe52c9ade5dda24 -a0050a819ed21342aac9d25194778beb3145f56a66980f620998923521ea -3f957b6ed0c5470734af9f416a16427dd03eff9a0e023452097d4ef936d5 -49a90823cef6de340a1ee02a52851b310cbcf41ae274947a62f9d1d8702a -669023e3caf967204a340694b45fecbda4bf9552f6bdc62d43b3b2c3d571 -9983c182453e22ee34241ab908e667115f7988174684cd70084aefc55caa -f5352a88e9dac45d1ea0e032af61fe9a9118a3931b2050fc6db66ab96a39 -74353b597f34dfd9f72150de23285eda5e555a607d198c291965a7233715 -3f4946a57af0b440ff8567b01a6f46c6d32fea5f8bf57d89dccbab7da882 -ee6c9260e89443b1d7db099477492bd0468850df3db668d741123e7ebe3d -c21748ab4c5cbeb5de33b8963aecafe76bba0c4f6ed8e8263a116ed85e58 -fb71ec4ab0071301be7c7d3afd5fa6ad46c0232807bb7fe129e44bfd16e9 -fd0c8bb5e7cdd86a78b5fb0669093c22eda9151d85b6f58a9c8ead3727c0 -09850bd31a8b4a873d0a506240bb2aeccb8dcb6369532f21d9b967aa8443 -fd6d77cb2d65c4678a5fad188db85940f0a187aa1031dcf5b8e0d0cbfb6d -b3b96fedec5b249b7a69de9b42dfa605bd622de7a220cce9b66e9f3394d6 -13487dc5e82c1e619079cd057b1e19ac05ebdfd7c8bf01c6c66fab49e0b6 -613df9e42beae2f7b9407a2bff8896d8035cea0fd5c11bc5889cb3d90876 -61766138d2625f42d0244adca65d1bc73989328c0eea0b97c7c766285ab3 -351ce2b183f774488a8806c33178090a3808f0ce5e339b87cf7add933301 -ca486742831ca751f0626864ce13172829a8419af5c78794a0eaa17b5bcd -fcb684f7d4bb7af15deb432e44dc7dedf56eb8bea08b46f1e8123a49a349 -a7cbccf833a528f5e22d2d463040e09b91e543a2f33077b3e7b9ecc64f14 -306186cdae1fc317a6ced7e9b4d51a10bbbcf2fadff876b4d9082e3f4aef -dfef230e4232572f4fa33a6e065f6895aa2ea96c5659cb579b023179f0fe -de7ba64bbd9362a7b2b8c4eaec254915629e81d01c839096339b99bc9e25 -84536955feaa52fa20666f65bafd9b2e69c3e8c15d24fa407e7d881679b1 -789a0e2a695d13553c92c0214c9b7562cd6a9a3d77c8b0c2196cef76dc51 -d855c1dac37f96eae4cc7bf07e17dc7c08333d7af33c8b2965ea1f23446b -3c96c52b30ea628ad572694d145b58a606f90b278290297aa372cff56b6f -56f4aad6612eb7c7bd07db4f7d1a70d8044d16d0b5c1605ee02a852ffdb4 -450147b3f9b87d72dc431b34fcdc899462dcc1b6bb6ab1758b6a589e91e5 -8f5196251d00133b43749b7a11fb67a22664c5e38e336dbdeb5509c2d9d6 -2642c07275949df0e2db59314ae0fb34641fc171d3fe1289f919136d853c -d9048ee9db50c699c49e27a8df199590bbc65b23b55bb387eed0c73f2db5 -1cb091f8c22af83103f214199e371f7de1df23f757817200be30610004df -81fe8ed6eba79e856fca21a126ca326ad2f313c16e15754663ad6a065e08 -4050ff005fc899d6e233691b918a093b5f1ffda8839ab23ae66b1bb7b953 -0a7f896ec55de6fb9faf1b49656ff2e57488cd7f1c44114c75f9d571461f -767a6040ffa14e9fb43096f164d60ca530d7cca76d526d1999ac1b52a793 -28651112a65db1f2564ecf90ea6bf2c9ecf515640719c3fb5e36cfc58591 -e227793f39b9d3a9025cb10f324a95c29c488724aa74812366ff0b118fc7 -19f9fd0f202a040be47ec99b46b4dfc3d2a17902a5779c8d52b27231a1bb -5cd794c838daddc3e6824ca8297ba669a818c239b389400faf17aa04b802 -f763029edb9784dfdc42f223e6496a938e613463bf9bbbd59d63300a9ad7 -4e71865cac4b4e81a5864388c3886e70799c8989188341f7d17cb514cd99 -3b211883f171ec6402cc361885f4f4b110757bb3e52941a94bfaebb2faa0 -3e32eb72e25e31abdde82c2a9015478afa0f434ae3f8b97a4bef598d6eda -44ffe1915c26ee0e8339d2d45a6a080550f538ded5542c8b96ca2f596979 -8bb6223e460e857516ab5a3323136ee8fc4b0556a7c39d0cf7acb45e48be -4ae9db325e4750b73289e36a61b301795bdb2ca2a8b933be1c09fd0cd2cb -8677df171d36ef1519a2269b21e4103b2ee151c513df3e10b2a216d6fb22 -18bf2005fa7e0f0563ad96661a7f55e1b5b991f8ca285651b2683c6a7c9d -2d1941374989b06f2e9b42a6af60193dc758dd8e9fcfc7c1aa06eab47e81 -bd79660666defac0c6b9e484df9c17a61ce7a61ef73150e8cd406af6da17 -4d9c2392cc420eddda40f975ffbeacad8ce1b4e14bee29ba8552ff03376f -c034784b38dc1d0ab7bc53943d2545b03d39797af8d58d6dffce56a353d9 -bebc833f04db321ca8642bbb7fcc63ed2349ffa08a33a5d0d78f4fd2c5ea -4258e4671e362036f1f67fcef9d878ae2c203fd9c05200c59cc98633e65a -99d912ec51d6f74500d5358b70e799a6817f59adfc43365d7bba1fd6766c -5c8e76248daf3f01e7a8950fe875d657397797a45e7f99a92887300b6806 -b86db61e03c4c09d6cf507800aeead874a94e6f665746752937214302045 -0b19cfa8db69230517183a03a16e5503882ea1e419c333d3e3b73cef6762 -873ac06bec34c3f736494483442619f5bbadd86f128a5a40b854051893ea -8d31dd6656777ad4ac2572d17c6fb21385b053495d1270e65d78334a4115 -2787ea89b86f97e72718905a11e9c5664837701a3c1c65ccaf26aebe8dab -c1207d5da2079c37883d9235708f370203b3b2a8ec3a5bb35fab93dae115 -aef626dc44b67ca56fac18caf1c22e6fbab93564829a75776630b9c42513 -721ca0fbb0b402f4d1db8f701d2b29fa60162feaa8a167eb3113c6f57036 -e8361357913eb24dd38dc6d3bf4c3176a07ffc75cecf8e5940a310f79a8e -f590844383d631796ade04a91144d073a9413cff34fb454f1fd75cfbe5e6 -525c3bd36ddab80138f6c19aad7417d47df1f1e0fc958fb190a8205b5321 -7c43a4dcb0599be404473d6faebe7240dc402a0e0caa21b56a601b154524 -f44988e5074c71ae8e1948bb2a2ce72fc24cf3b1813cf7408a6b097aff22 -f9d285134d09b7053464259531eb7b270cd5f39f81bbf41a36420f61e5f6 -b429036bbf20e27af1a437becd74c5bbc25ee2519402454fc94d430636e1 -736fe65a643d9b9d21c9a54eac5a8fed51ff60a47b85a0e9423e330e00cf -220c23e056d20aec2fca3e6bc7a61a8366eb940c9bc99fb90e8704e27655 -20335a983eccc7e20b13745c4b4f30a842f1ba64745718c152697c688c73 -6cffcf5cc8eb5756201560413117a45ad3d264291cd51404f98448d31474 -d47d17d201def12867ba679f0e2605de8f3e8135ed0234890cffa68848f0 -6de427741b34c2ea654251ae8450a152538eb806ace3ecfe86d8c4a137ec -c98c6d6cbdc191a5f8f5b5972c70b4896960037b6d4c7c63586a52d5eb59 -47af8c192eb980d0801fa670bb1d08740819f9da1dd9e153010bf9580a1d -0925d8327ea1b88db8d934f40266ddf93e5ea137f267847d826cd7999b33 -c795d0ac05abe2ec1770dd98eea67912f1939118defc9b379e237d6477bc -91ad08e0046b0836fafa1272b0213dce990c90815f5b30d0eb103ac9539c -2f7bd2280264cd95b4be84cbc5139a7628ed211905dcb92cbc3180ac9e6b -b9ecc3cb08608b2395827d5729781dea49d328ba0c1b4cf2cec9f6bbc822 -1f2bbbb9d88f9e7682b9ecc06b9705faa8a90a51678183db1e24cc2c4307 -e16b3c20f08f179ec69df7a8c4261427f5886f9179c493bf2d0ef36640d7 -79925585724aba69df6d1b4f0bd2a356eedfd74a88bea667b102420c2300 -ec420e99b9ce8be1472b617e1255a7f43a0b52f11657f1a4dbb624a24886 -9604fe2062b98f5787d010723e520a4f42a0c3943e151ee627f3d5db90e0 -7747e1a88a53c4784c8d2b042b9c23c9e436d7d88343171161a364cd8961 -37a19582a00d774ef01c7c3fc9e9c7be5074c858d2bacd707a6a4f322027 -137d6ca0421ed9f9c7e7229e867678e5272cfc7156a419e893404ad7dabf -a5d8b6fd0787cb4fe1a901c34dd931f1b64f0c470ff807005fb66350d0ea -eb84ebef2c2399cd14a4454ea5004bddd99988b39c4134b92121ec77faee -55cc716eecc58b594b39c41dcab308efa4458ed71943ec5805dcd0194ddc -1ba04a5d3d42d07ac62a907ea25cd2a7e77aba470324d41dc1a3fe088388 -787b3312f472cb4f23a414fa5f7c7a0cc5d121d7642b5b3f0cf7ca2173af -3f878f374938251feb3ce5ddd2d7703fc79a130978ac516daf70ae903799 -28bea3a4296f48725d578d2e8fb0f932e398404fa8a242024bc011c0ae81 -7b92bb104712253a5d89c543a744332069e33ca08bd133211d233ef799f2 -fed6a20a9073021e505def8b79e1279dacc062cfd4dddc2e8e0a7fda5dd6 -bb5a745f99cccb7ec1df532308da3da0f236c74639c280ea649b2f7ec27d -24221470b642567f3b2e1cd0b3ffa65c5ac986b557aa9b444bf470380435 -abae9b51c6da7ff753810ca7938d8a1c47d2b41fafd236cb5998f3ef365e -1f700bb257679ba3a82e235a3e97a667a6ad94412839c96dcd49dd86ccbb -6df8ad01756b311e9fd57ccd2eb2f19f035e214804e2b77769319a5389c2 -35f3ca2a73c616c9ef0984abcba167d7d652b330c68f4f6378aba69628b4 -2d59eaa2a7e4c782f6eb96f6758d17d35650b15cb5de9bf973b3b6f67c1d -f3285be8322fc2b44359640a3ba5d6d7b96142583a00a9a0ef84fbf14046 -09ad55b2aefe8c5c8f58ed21623bf765f81dbb6cca6d2a51fb7730a14839 -392cad6b47f5e03448350ab36a37d9ff2b9dab69be5196511072b10cc91f -2e6b5160b2b1bd112e6c02d14063a9bb46977b0d4bc79b921fd942f916c9 -c5708e0d133c8309de2f6ee0b1afc996c889c36de20fbbbfd32878f477cd -7735c7c3fa59e9c46e654ea20b4381d9f6c6431082e6918d532bcd539284 -af0333a783c9e7fd4fa1e4da5ce8fea2ea4037644a24532d65fa5c1ee982 -89e4b9abaf71a35d308a9b8c337f70babc5fc8dbb0327143707ca5b675c5 -2d3cf09f7a4f667fcda03d8c82d157e661517787ce6bfb35ea772de13c66 -2bd24b74ff9ab0fbcf6635d8e06b54b5b3125d17ae13d175cb7922338ec8 -9d1159fea2110995ce48f7d2b094f06d11d59b3a64a44a83d48c78855e47 -21243e82d9858401b094a236fa0a90d61863931c30d13b9bf33a35ac0d11 -a999f2b4dfba6fc187f8c235a5217d777a5a97112e7db6a8a4b06b07d9c9 -f41820e233c8b58b9e47ac56ad1ddcc0b35dd03976bc776c6ac3692ec0ca -f8c75ea7825bc84156468ca7b269d890ec9d4a365b0b31d2f6530185d5e0 -2acc3ce14eea55ebb5667067825a8682e135d23c78863d32065ddcf1a755 -e0de6dea7220d1a28416b96db40b1e9f159aeb070c9a9515f301f162b0cf -e32c6c89287de6e2b40458e3393826189a10af8517ff5a10c41c9d05d999 -aa9305a2ee8e7fe46076bc9c5722ee0a140a144ae383e84a8abe70af5d29 -96a0a896cd499caa0ed7867e7c3aac563763216e7769d12218b584d853ec -01db93ca22d0c8d6b286b20b6b26d6ef19f2cebe7030ecaa68d069fac7a0 -09d61770b5e8f83024a99142f59d88297cb8d093992c3c6c11b043b151e8 -20df640407d8bc829bfc196bf2901e63c6f16102d03ffb7c54a7a560f5f9 -5cf8379f4a2eccdcb604bd553e6157b4381940d1b3c768dbfbf2618812f5 -7fbe744b3d8ad680dd9223d8bf2412ecbb614d05b485e3b4669d22b417f5 -02cce2d705c208b15fa83b5be77ccfc1c840f385a58ae49fbe6ab4e53912 -473630e0cfecefab95ebc632a2b10a2103bfe801ca0302542080cfb4cf4d -4c241b1a6c8d28114516e3f1bf39dc02db73e6d9a797279acfd79b02a71b -ae34860dd0e11b18954129f8dd57c039bb7063a4c92f0f6a1e25f4ae59d6 -6c1cc6b73a79d6a56f7f2a8a64d571caa8a760f4f485d770d000ddf393ba -784bb27b781c47678dd78ae9b5d5e8b57d163c42c7a55e4aae22061686bf -aebcede728ff2f65e75955585208c176d100912836b5200a79062d4f09b1 -ba9465b0e937e289160ec543a4cedbbe0cdb5ecfbb4838138ee9e1ac757d -3c5f04fb6b510b389e2f521759e403bfc8ec6bd79e2d40bdd81901c10dd7 -4620acaac9108940daf03af23f09d3c8b785db562b05e597056406557857 -e96fc8bea53c2c2ccd0ea6572abb0acacfe29e737173d665ab6dc2995f60 -807aaa4073a183aed23c26c67eb137c937999fafc63b66a021125e4ee5c1 -a745ad1fff2bd828dcef392052965ce0e9af7a2c88d730fef69da91083fd -83d9fe9f73d42a8dbdcaba85b0fa93b210dbf49cdcbf5d4b69e07375fab1 -a39038cc51f66f0b10eebe0cc61f697f7025d9755830b2d65f1ad0db91ef -ebbfb578053de329935bb28d6ed6c12f748a2f70458990f04d56c35557e3 -8bc5d2e5de7f52bcf00c3bcce091aaa8852d53ac686f8f407baf3f7c8968 -69f3b62f44a5e2291aff9d30d7b5c663658a41add74562dbb0f1062f564a -9b907846291700151de04c1a55cb945eaa2e7a709218ec56d1becce1c0b7 -dc41d5f016ae8080c3b07311590a0def35337fc3c844c0ccd04926be9fec -509b1255ef12f368d20601b1ac8c68b0a935f987a21de0f8191604e921ea -0c04b00dc188fd73499852dbcccd4119ef799472b353be7f7dcc904ddfdb -920839f3d4a13bb1796f2dc886f31217845f8d7a543aabbc720311fd0e6d -a31ad3daa06d5e7e6270a34304f35ef170a7abe733428e96b0522fddbb5d -eb35aacec147067fe066c9ef145246fa3d444d176c274b91fddb8a7bd7ff -7cc7693c25895bf931eb321dc9d79f662a17691f9bd1662fecbcecf6d1f9 -cd8ddcda56d19811f05fa48bcb492feb355b0ec7c04d6046549c56f7799c -2cd0d9dade8809de7d510702e525ad9cc82c41b4fb36218e3d72e905c507 -159076a9c0e4a008ccca17bd594c69f5eee656426f865fc1988d677b72ce -b710b29a0aa8f8337552ae30e93bf7c6e5d013555872dba4737dc5f08c0f -efd428c66fc8da675373f13f89102688977e18e14dedd7f3b676256b0263 -b66b013617d9a026794b0d6040c23c5506a98530249633a6beec46117c96 -ec036eaf6439e25b8e57754af5ebaaf9b57880ad4fc93f002fb03e9fda21 -df4acb78296b0c49a5a852c134c3b10755177a0dbd6c54ea7a2b9bdac62b -5d7f3da649df856478e4baf97899e0f891a96536c283f5c81200c51c6ab6 -77285450c7f7e96836b6da5660f6cb76782ddfc64b6fc348ebc3ba4a46f7 -19176296d8c5a31132b3fa7d935a5d777c1dc84d669d564cb4fd689a38ce -680d0b3b130caea0be43864826d0d154019fd0d865f1c389cd367cb5248e -24640eb6f66603e50581f6fb5aca6cfec1d6dbf4196da10a5e1ebb14e4ca -0251c4c8412cc1673d6e7a9666b04b090567efa0b830d2362fd384cb0303 -8a40290597bdaffe429bb89fb66b9dfcfa92f39d92a8baba7266d144ac04 -f069093ebb3fcea961ba4497d3628ad207e0c8c4fac0e5f3f2a663a8d05d -b6dc33b890ae13d84dce64b495d24cc749b121659373ca31cee09bff2e9e -e5b62e89d5faa4482a75f341dd172500a54b98fc108a69a3ea94db696513 -d4c7691e0095ed3900cd4489ab008b5460b34ae8dedf3721c60de7086605 -6c391137cf23255c565bf11403bdeecf8bf39ad5e4317a4bb37003b2e7c1 -400c3b8ed7f63719bddf07908dc2decdb0f68e8ef722851c4420303f6de1 -b5efc9b2598732fd1f2cbe45a504bd7fbfdafeade3add7274a1e875aba3c -4e0abfc6444944b79f95b5009560818f7a0599e5bab4405378fadfe084f1 -653e5a0166714047e8bd4e4cb116596d8089bae9147ec1d62cd94491af75 -a1743d58bafa11b63b447c954a8d7fe11d39d969feac8fa93c614f97807d -ac62cb7a84a974a0fa555a2e3f0ef662706efcb828ef72e2ea83b29e212d -f89ffecabcb08dbb7119203c4c5db823bf4e8b698b763fbd4d21e57940d9 -1754959d21f3f649d856ac6615eac692ebcbac555f772eb6ba3cece5ebfb -cfcc2f3d8dcad7edc697df93aef762cd47cc3ba9e2cdd10940be676efe7a -a3749170edb47b7562805e3f8bd978b18057c9110ff8d19b466ea238af32 -993e2d3021745b238021f824d887d2e01a7ff12fc6f084b35292f4864579 -406c0f61d0ac7cdf7e4770b424e2ccc22353e6c82bf8ff172973df267ded -bdaabc2a742beea02e35b9b253f98de9ca131f802deee2905ca1a6dc4608 -19a59b4a4265c723007d0215fc8ac2a91ec5f86cd6aac1e370a297103c3a -3cff58c7ae201cbaaa8a12c93e95e73974f9abcd678451b1db02ebb2e10c -c5abfa573a2ea4219fd1851765649318bb556b728d432ec05a86e9894aad -9cdca63d08642655801bb37f28b6e11b958e8e800c8d521ca4aa045fe9ab -ac02dc015d18b1901d519181ef60227170a07f3328a6d5fe4c5aedb35fc1 -3dbe86564a9b1dd4c7ec648880360cdd1742ed4ac409450f1d9681cb5e46 -5edd1de2a2c7f8ed63436f98e849504ae71bb872683ae107ad5df3ca0b47 -a5b79513e02d7c540257d465ae4521cb3449d79c931e2ce8c5b0a0a4ac88 -cef7b9e5f92bf721ad51682d6b6f6c14747f78eaac1891fe29aed4eaf177 -e3d2fc655ae889c0c30a3575a76c52e95db2f6a4d8ffee9518391954b92d -39dae4e97c4022031f8ab390b66ada6dc9ab2de4d1dddf26ac4032981a69 -08f73d34b4849ae28832cddc0dcd116a47d9262b0f93c24fbfdf8a78e6ae -ae3357f3fb89530854257a9db773a1acf5271fc4ca04a06b46dbe661ca11 -9f45e0080cd129e1a7c23a33f1c48af960761b117d9d91fa5a0ed3e47865 -b774a322f7dddfda2960b91fa7ba20c8f9eb213251299ae328b28ef54b0f -55fd54f8047c555e4045cbd70964e1c953e471408e4f25fe8ca7009bfe44 -0244b1e30dff518ea7ce5078027baba4e07ecf0ebecb497b4bd88f1ff72e -b261f6dffec0ed895e237b5608d31ef479e8c9ae9003039a5fe67252ee39 -774e1501100c0fcf154f5c5c81c70539e03118ab91f4ce247f6132d46346 -bbbb126c09d7459c1977e6e367a0c83d14edf7dea081e5f795a7c831fd1b -325b33674ec9c2b68029a0e600746329ea2e1b9bdd5cb2b140468e53c108 -8e8f2567425443f8146ec37101fa4dfccb0e032fff6cdfd76382463551b1 -ae8ca6cbff0e34a3f75ad400a9573217f8cbb00a6d59ff46e48421e97091 -cb17f53f20ebeb89609ea55ed6ba4101f2f3ceccbc7ade21202439ef91d8 -a9a783c22de7e6601b50c4342e094d0eff223494489fa92150425da1b432 -908423fb3f41e0b115ec1ba592a4f920d15610b9fb33f9912aba67912d05 -1ee00a13282c1909a3a56c4ed06f2f4d1739dc296b7492aad0446f87a416 -c6db4d42b504dec3a6756f3d0845ab2d2e151aa5fde12b31a9c3b5ae1cc9 -d97192bc048f00dead66940004281c4d5a92c20b1f77795cb4f98b8eaa7c -be16f9b9d4a34a1a53e0a0deadb4fb4b20d9e8064d3412ea8d2ebd259b8f -2f04bf4bf11a5ab7883c99943d762549c3d5866bb6ed85a0e862eafbcfc7 -03bf4b77cecc0d65bce4df33e0d65456397f231f8cbf66672457cf539817 -6aa5292fae24695009e55904a04588659a3a23fa11989b925705ab45f954 -6f862b0e176fddf75b70d9ef7389f750becbffae25d58a1252cc04a79e13 -fbb6a666fd87cec5562c3e14fd78ad05be28ff3871d6fceff5aa8965bb65 -67ec76d105a6348e915b27767f5010011e80e0e2f9c34742a4eeba369e66 -8faf086a45ac9bcdd76c758db01a78602412a4244c759ece0b963d9ea58b -0efbf4376bf115288803a54cfcf78584c8af80da2a3324096463e3898285 -57de6c6354444b12a74d5e66053f6907c48522cae9e93bccdb4632131add -52eb374213888125de71994c31dba481b70b2e4c1f10b865d58ef09fc9dd -2ca7f69bd2855895256caa5dd6bf7d4d8b341d677c56ca08fd7ba37485b1 -444af8be0dcdb233a512088936ab4d7fc8c03139df396b7408747b142782 -d9406db0dcd31368d2f23ddef61b0da3c0704e9049ccf7f904548c3ca963 -76eadf1ccf77f94c157f5b84f74b0c43466134876a90c5fdc2c53af70c3f -f5c2d13cb665fed9016454bac1a629361c8ea62f4b2399233e8587db6e75 -a9cde3530f20a68ec155d275a4aa6f63aa5cd115244643b54911c954feca -d57be2a6c40f1bac38e393969617b066f7d94e8b18dd80fccd0168d4a385 -f2f1489d1dd41b68d47e5ec66ec568333d1f584e3dca90f1367a990630d0 -14355be7dc45378aa111c319838edd441f15e125f928e044640f25ffdcc5 -c116c3f6ce0d4d3195187b22200808366eca9b508ec45e664e562186efec -a97b22835d384758849605a01973cd9ffc1657b124950c9d9fa3e18b1a20 -7156c4f96f08b87824373c2865845d17a0dda71b1d69f5331c5676d0648b -ca80a7958a2aa034d7e1e9fafead9248e6e64f9ec327c60ae4f724e1fb95 -8a71e82ac3842768b27b506b5982311557432dc3f270ae6eab23a42fef70 -dd0d407a02cbadeb7b8b74a2523cf46a5f61e52b053c2007f75ae053a96d -e00646662d027d93f950e516cddff40501c76cd0d7cf76c66b7bcd1998d2 -7a19f52635c8e27511324aabbb641dd524d11d48a946937b7fa0d89a5dbc -4b582d921811b3fd84c2a432dacb67d684a77ac08845e078e2417c7d9e08 -bd555c5265024aeb55fef4579b46f8c5e79770432c5349d5a65a47ce9338 -e1b599328bb1dff2a838f732852f3debf4bb9b828f9274d03d7cf813b123 -687c5e78a26310d87870bfcb0a76bf32aa20e46f6b2826912e562f503aed -11e427b7765cd2a68da2ec0609259ff14f57c07963d075e96f8bd2eab9a0 -dc32714dd8905f2627c6d6f33563436bda2d7fa9a976f88947b84c72f454 -bf0b66ca84470375d2ff252b4a2df52ab613d0c8ef0465ff1d809ca82025 -c2122a8f44c56ebfa25690bf6a05675ebb8634ddfd24c3734fe8cb32d6d6 -c69c72a4951cb959175770b4286d383e7a3f158450945c8a2ccf7e54fb19 -aa8d2d98a07f0c55f834f2728d89f82a598269750115a02287c4d415cdaa -14e1d9e7032684002f90603c0108dd26b40fb569bb21cc63d0da7e9e1873 -9df0a9c85bc340d2b0940860d95571dc244628c59bab449f057e409e58ca -cc3369f4baa8e53c6765a55620e78341dae06e5cdf2fa5e5ba58634b29ee -ddfee7f78672e55f18a7debbc30862f278f83f4cc123ab591371f548fbf9 -bd24b3453b9b57051c2e67edff2104f3a05a9f0cb7efd81c1b1b0a2bbe95 -21854902526e5d4fa1b3be270811b972e8726623410cec7911c07f871428 -1caaead97c503714eaadb14ae5923f020093722df1b9d9c055d7d5f95af2 -a9fbc5ab6f6c2bd655f685534d7dc5fbb5ebded6ccdcf369bd83c644dc62 -84c2810495888e9d8f464a42228cdc231d5b561c6b210bc493fc1e7bfd66 -5a6c4055a6a629f571f4f05c15cb2104b4f9d0bd1b1f0ab8252da384eeae -f5fd5c663ad7a2c29f65a48a30ed8de196f9eb8ea314c6e86989298146a5 -589f76f12664c8d008228b33144679d16ff564453b5e4e9f813191b6c99e -2680e20a410949ac30691b1428a255b6185b7e3802e8511192e73c376f3d -eb807ad2727fbb4b27538b3213da0746231b1c1b595a958466155835c537 -e0df4a0ef272d4c3f7f2ef011daed38bc58bb0fd7458e48060db98971bd4 -b24bc7bd0de92573a1c7a80a5fa2b34fbe50271dabeb83aaa4235cb7f63d -6a6b399360df8b1235e4e9ab59698930044a98d5e083b5f5a5772309b390 -9e1ff2a252734b32fee3940f0e1ba61f54dd1d3f6ff0d57c9ae75a302d14 -b9dd9034279aaca80b6bd05c74bf3d968305a5046910871223a3ef8c77d8 -25d7e6d3d2809e76064c473d1cd7c05666040b6eba647e34588f49fd70a0 -3c937933a2272c938d2fd3aa8149f215bb48f3bb45090bcb9a6ace393a44 -f1a9bda2ad09a5f566b2e8887880afa45a603a63ffe7c188e3eae926a903 -4f1803368e773f42c7391dff1b9ce8599161515c549aca46aebae7db23ec -8f09db0e0f590aab75e8eb890df354b37cd886bdc230369783a4f22ab51e -0f623738681b0d3f0099c925b93bbb56411205d63f6c05647b3e460ab354 -1bf98c59f7f6c2ea8f29d8fe08df254d8a16aab686baf6856c4fed3ec96b -0328738183dbc1eebb2a3d301b0390ed8bd128bd8e7801c89941485c3c86 -22b5f223cb07dca74f0e8643240044e8c376abbd8c82ff98c6dba9b6d244 -5b6cf4189d63c6acd6e45f07485a0fa55eff370da7e71c26469740a68627 -a3c297d2bf215121fb67815b7b9403aecca10d21e59fabcbe38f5ca66e7b -551b22e28f2d1fd7303d15a42c45bf54b40ef7fc93060ae5164e54f91c55 -20bd303a98d0667a02a900813b260c0343021ac01872fd62cb6abebc7ad3 -a4456805159839ca4a3e35db586221169ded66f852e8974e3815d4d7659f -6a9bb93585aaf264f06cb6da6a26e51683945224158ea69719b8e4e36eb1 -01333aac974db8f84b051724cf245fe7a4c86582b5dbb9a5d9318180e33b -8d92c22c44b0d18f8ca34dfa4ee9693c1a26fedece01635fc5eac1fefa81 -32458254ad46dfdfd2be12a1e7f32f3728f286f1d5d4394424a073696b65 -e3c459aee9310752231fa703faf35e11796c4eeef698f4109ca8c46ee322 -5dc2e3e04fa787188e583321f8410b68b9624ff60679d3f25c13e5ea7506 -a3ce8d0bebb99d9a959ad92d8cf909988d9250b310629903d6bfcad4581a -504b91b2c91889987f36d6fd0be1d0ee5aac00aa0cb48d78a1f7a64a777f -089573ba79452efcc31c8258fb317369feb0d7ccd48cf13da6d1ccb59a4a -48ea0b398e590c1169113fed81639e13e96aa268d99cfdb7aee977fbe85f -f784853a06642b5521ae0a7f610c9739af31ba7a5157ebbbad999e23794a -d2cf25af987dc85dfa29639957cf28e7f2b7671188045130a6e2785f8d8e -30e91f0f68c1cc9f2de902952730003e816e4f5703db7a97b4c566f80547 -42fa77be563ef681a4513b9a68b2b0956551c74545cc9883428dfa72fd5c -4eee93256b26bc86ea34f7427cb0c0cc22c0cc343f739c6c0c46d0923675 -5e04d70587426ef875f8c89ff8492ea23e4e4d763b84a6437a440e69eb70 -65ab6d8cf5f8444a844e6ef3d158b451d121daea2d0e2b423eea24254226 -7eff1b4224c4e80af2a7becac1649e4bbef09f39415e9b1e3750d7ac47a1 -068a4f5ce30840b00574eb4e683e3ec25f6e690feeb0d354568efbc354ba -813ca1400734a67693af127b0f636d58b83e91548f98e3d87da7fd7cdebf -f3ecb4b9272d1c83d4980170378d32f1d98b87c440881af9ec052510982a -0c02ba6743bdc7691a44bae5e044c25304c1a2525cf2c0694494a2e9aa34 -f36af43ab288807ffa4bd418ad51d98c75f2b2f01abfd834d3305682b6b8 -62ef69d05962aac485bb4f560583a5dbb74e967eaf6d299160753ec32249 -bb1d9851d5441cb0c624208e69dc876cd8841a66976b5d7f9c99be68363b -8112d33d971f2c4f2a1feca88ba1a794ddb725c5e2e2c248082231059aef -729bb5fee5006ab8809f63e162fc0743c047c7984a9e6333b433fa143d73 -72d4a74fe37314508e04f54dc7a1445e2d6178ec9c041d0cd4fda5cae830 -4b16feb21f3222261c293a8b058dc708405c1a97ff34eee4ca69ff4e1ee2 -a03380d52297574e3aa50c8afb826fc94a14e8caa9ba89d6e92913be9e07 -bf7ae011e6bd142d8952d9c2304735e875d1ddcf82fa9fc0c6449df2acf0 -d5f6cff6d21ef6b2d29022ed79c4226c97f163284f2311cf34d5b0524a1a -a446645b9d05554f8b49075075f0734b3d1ea31410759c174fcc7305d2c1 -d7128781043cba326251a3375784a506cf32d6a11a4876f85ffa2606fbdf -27dd16d64b2108d808e33c409dd33f6e0c6079e47e7196016f261e824fba -b0e4f91a189747053e648ad2d942ece8f582f052668b63a23a2fae4c75a5 -180db7811aac654270ec6e341126e3561429f1d41fe7ba3f1de9f8bbb8d9 -fc5cebdef869376a2e42dcaa578c0807835e58d75c39f91a83d5c1eb86a1 -b0f7aab991f65eef030f212d38d10b1913bff71717c06c78d9a1be136f21 -4be157ba11ba309326c55c23ae8512646751fb82ae200c06bd2e644bed38 -c7cee826cb587ee8ff378b7fdc00ec316bd4a9c24e2c250cb3d64f8ecbb8 -7f4d81626d7f1e4491908bf17c48c84bb1736693eb4d0fe634484cdd590f -a40ae94d44f348ba683a43004b487f047745fcdfdee2e913328a11a99530 -9bd117e0e5be4fb25d176d59dc2b1842418141190ed9ae1f33e5354cacfd -a5e4bc186119e1461bcd98517e675276ddf0296d3b3cef617dfa36b4759c -944fd721e1bf63d45cea90b5817a40d153a2f779e03487cad3c1375425ac -8cbabf7f754d16cabe45c65f1be4441908e0969d5a5111c931e724537dea -7cd3fbfec9b2f7d3efa747bf586e9218c3106c49276b89fa28f770fa0644 -fe1f3fe3adf07f59c755a5b39a2ac1d6f23c256a293bf3b31b6b9cf4c622 -b188d6e7401c038657c78bfde9ba09f508f1bbe3ed79793772cfc928c4da -519f7dbf3ff7074284437d2de8d7b7c78829642d924abacf353119e9088d -14739935a23667c432806085c3af71ffb7c5fe6b4412b9b1044c1e62ee0a -a5ce7e0322bc65a8c7d874270d84136526e52d0c7f9f93199c6bb7301216 -a19bebcef3c5633f21d012b448d367157ad928e21f8e471e46982bc46a7f -df1bf816a86dc62657c4ebf286134b327ce363ab6a66634eaa2a42e99034 -069fe1302febf06959eab8e7304da4d94a83ac1650a02c38c1c4b7e65c43 -e3a6fb0213e57ac49e58721a4f36996069caedefeb48f1a59303459d5873 -f3bedcdb9d00c1cf31130c27b60928f210e1aa5e1c8e04b86d2049f31265 -9198fa646c53afa9058eb8ceb41bda65f415c79ac92af5790b176de1d300 -f1c06b782d584f458dbd07d32c427d894f84215a8e7819e295ee98d976d5 -644f11920ff2f49cb1075c3bb42b9fe4b561362902f11a75669b7e7c4475 -b65f1ae48834cd67816eb63b58cda2f50bc22eeb0cc965569b476bedded1 -2701668f609393659b266bb0e37bb27afc90bca271366e34754383363592 -0f9a3b508aabfe8deef585b07a992460c592a150b325b1e50e4214a2f483 -e9dfc826c54b488493a96eaa37276f5a9666f0a5388fe388263d2c0cf614 -c6cd01571da4389f01fcdbd0ade1c435d64c5921b5bf7dbebd5268100a03 -1e1abb8cbd83873089a9e08cf80276c7e30d2bb40280278c29fa818eb079 -87623b1cfe13e0b01e27be0a8320b69b5afee820f4705202158b7f3059b3 -655bc28a754d088fde23d43d6a9389da8bc1cf3e8ea1a6f4328c196e655e -42184444d8c0614c7167c91a492c24c8357794c61f5e47cdaf4b38004a5c -8fceaa8151e929328bce1b8f67b22034f3f75e4d105283337c3d460e7d99 -89920c43f5e1449c74ad6ab5ea029cc6e497ea60068451c4ef2132fb87ae -049077a156c868b768df4a4c475a532e2a22d999931c64f8bcc18f51d25f -0f94fbd3e9e6c094f78da062f80c4aa2b86fa572cc469e629deb4ba0c553 -55e8422b562ed2f694d0e8e5540144e30841d7593b255edd4a61dd345d5a -00e411d2c50d64782a3ebedf945fc31c00d2fe4ca800f5aeeaf12ab399db -956362e979bd7ef0787188e43835e5389ac444d13204af6bf1875622f175 -09f32015c28729cfa3b3cca90308eefaf260e3fd9df10f3e76786b8bc0eb -a30e8cd33689aabc55e3ce387cdb89a30573495852a48009cb58a0fd34bd -da911159ccacc94698ffb94c5f45f15ecc9e82365174cefbe746f95eee44 -7a33b4d823487e203478eeb2d8c4bc7b743427778249c56e48fe17d0a501 -7b693509ddfe1f42bdef97aedcc26ceffa9357dd985cdf2c70bbfc987354 -6f0aa7df227ec42f9ca2482f58809e3f9650444568c54d3520bd0a7301ef -48bfebef1fc4332b5ca851fd786c1ece136fe9e575b69393b5aec2611903 -fae6e7a5046e2ff350becb8700f209b1131044afd32fed1bc1297b6a2f29 -6ec3b87f170e92aabacc8867360e4dbce9ea29f0c1df981f6cecc8986767 -0ccfb4c9faeaad7ca9029b8ff0129fec4a040f80ead041b3bc8af7526675 -ed9e13204e64d76440a097d77c535d34165bfe9ffcade530abcc75ae224e -890d5c110004e218bd827a02ac7340e18bf3684c43e664e0a37d5fd4fd1c -4d4489d25a99d542c16e06685652cfa3567da4eb0cb517be1482939da0cd -d0ea3519ad1e51bd9dc7b9077375a8cd3b5de9888697e853bacddbbdd1a3 -0e442e1d6f2d652046821813d0cc0e8f16c97cdd32daf239f5b2b65ef620 -46f6e9821b2e2ec539302747795fa746318514d38bdf0d0e490c00e114d5 -03e7fc9a8fb83b14337a5bb4d640b52630f5450bb3bfcf7cecfbb1ef5192 -ae401265450db197bcfa07315ff95a809bc5fb4249e3a728a817f2580ae3 -50d8d6577f79c883ab4a3119d9ab98219aed0d1e826023a66da814396058 -d95e52d9af8bdbcb0454721f27855b686d13bdb473f650c9865f3e04f08d -b10f5256a3e59bcf16b12a84bb7ef3b370647cdad5929b722a05f5b3669e -14c232bb82fcb9c1dd8155ff4515f4e83c895cafb86754e896f38e5f3beb -5d29f1bd99cb8a09c5e50f412f6d8a773b79021ab2c4831aa663c5defc4d -553616874dd5bd8b75c7a2af7d029aab5a72528fbc4b5ee3d30d523412c9 -60b432434017c4cd68b2062d28f307fc287e11663511d1a6b52143afac0d -ce0f7ba3f326fb707fb8d2c985dd60090e6664f2344e098a7a1a6448026a -2ee651e8141cd7786b6543f512e4c31d25dcaf6652b1eb52706300b771cc -0c49295067befc044ea46341927123ad4b7d094784bda7fa7b568853d0b6 -1e4cc39e1abcc9479f91a2501009ae34ef7d5ff56205cf5288503591cc55 -c48abcc78daa4804549562afc713a4c11152e6e4331619b2e474a25ffb62 -7c46112fa4259f07871f8d6882e9a7ec62d20a86a0c502815d0a8f3f5ce7 -cb4a6a74b6db8e17d54bc919b82c7c729cc05b98855b9d8a0fabd8a9bdfd -4333f395607631f57c0473be0fb290c4f40a7aa6ac49208570ffa1d0f849 -d4871ebcf9ef6f5106301cf54ff8cc9918d6de74d519fccba58bb1c21543 -f3bca9f43c211b2e5c233ff6dff2c9b56d3f656f6070d13dfd0be04653e4 -98c670770e01c07b731ca0e2eb56e608828fedaf1a31087f2d43cb4c0074 -e576769b0830577c86ad5de48ee216df02d7c4e4ec231afd8e76c608fc9d -06cc86f38cf4d839e0a0829902f56cf2f86f08b975a6bdd0642d6b4c78e2 -57cf9a4f52646a952f6a220c36c91db7f44c7f44bddf33328ea8cc01827b -5f2d79e3ee6c514a4f8597a847ef5f32c6400736e6ade28faa7bc6e9c6ba -e4bbff236fa6dd2b0ed23fc77f92649feba149f82488260b0bea2a4fe1f4 -65d96d8c51719e5e10d4c17d1b67e700aac36b1ed55c93b4b2604e72f51e -b30fbf5b64c6fcaaef764639ebd789f82ed354712c7f9fcd1df257e14c0e -8fd59a0eddab684bb1b4176d79b22ad2605bf534e4b8fac2272fbdeaf210 -0424a2c5cc65f8dd5faa13313dd926128ed466046ee94bd3eb41f3ea5505 -5a70603a2ae1981bfae8e77d850fc5a5bf1bacb3df9b7cbce68ce7979fad -a73c2900526b68236c6d37197b0c521c5b1cf5cbbc89238586eceb99818e -aa47ca94ff615233575fe83d0d50d734351e0363030a12300f7b20450946 -17bb209c346ac1d35402b617d6260fce04ce8b3231ab5c05af30b0f3ccb3 -3616d3df334c8d963279537563222dfbb705c3e14616ad01927f952e6364 -4c4b7fa44ac97616c1521facd066aa33b2296dc03682eb6a3b9dd8e5bf62 -53f10667ecb07bbd50553f1b211067f5cf098b64b84d94ba9ad8b146dc9e -8e9be06bc14cfe0945e22fd819856d6996e857c0bb5f292defeb493589f4 -515700753885d61eee1b8c19e6e94fe2302c07933f949d6bf119d207fb04 -dae7bcff7578bf33d77e29611c7cf03b2df12c242827ec4c4e5b5343ca3e -4f7f38ed337583e30dedd78a082f41d60cbad55d59dbba11af1bd296ed6f -e31d2e10d3a8b5ea698e656ff97755a47ddd862d23309e2e6ed3e3e111c0 -2c3a713d782fe301dbaff0a4225f932576622d1cbae40d20f46958298d01 -783851c894f2712bfc4736d3802e548a704878e2d139348671fb96d0ddbb -f56d9349172caef0dfed4b84d867116d91063dcdf9ec401dfe8abb269ee6 -0d646bd12e0752313e2ddc272d9f4aeb9d940987596ab623f9198765cec4 -62f7b6c540c9a70c9a872bd28ea62e056560b61ec51fc68eafe008f20760 -246e06374ae5a6bd2577217700507978811ec29985ab644e474e41e8a105 -295fa67ae05e0739e8c7fbc51104522934942f53e1e1df1ec2a66f0a74b5 -9885cf2c2fad1cab3e2b609f126ac8b7350d5408a7df9ed5c27a10ef6505 -6f0d877cd7bb902977ba93e6e8520d2d018560ec8143876ad0dcb95b173d -af72c0d413bbb5541f14faa57eedb3ac2430e36911d2f486d9ebf9cb6745 -2ccc763e1e46e7a4b8373e06082176a6c66d045e18f90b4b2ad15802f6ef -cf2130cdc627601ecc19887784b6de7fb6a193bc3d057ace29f74199acae -69526ba6f7a2c669593f9d0849f12e37201c32c88384e4548a6718cbb2ab -714ccc917d93b865ac7d7d4dbd13979843f4f5c1f8b937ef12fcdc9aff50 -f09d2625f4367ee70a98772a273d8919952102aa03297e3cbcd876da5abd -2ceb162b8fe1d9a22ff694495528c09a8819fbfb6946ab205d4b2424f6d5 -6fa1c704065cb64fb2aa0fdf291fd5e7daa38667e6d8e889be7f4c453da0 -59c492cd25fcf4a03a6995897145273a66cd6ba999138bc8e2aa7d080f9d -231497ed28a9a27b6b0d4785bfaee46fee71b26d6839f2549a14e7ab7347 -0b6cf368d2d49e74c78d93477828e4582589cb447d795181d3f13dd8ad52 -3c750df8f19b3260c17a6598b406472a7204dd26c5988911ce9884de9a1d -ce33d834becb1dc80efb07f32d3ed6c2a484c5d53746071576c3f67f25ff -1558986fe2dc2265b4fff79c07e3f4c6c0ce8319e04c14728ed722cf214f -65066148bc817753dfdcc0950bf80dc515002e1a92e7d8936e9b3aa9635a -a6d512c68aebc79a62a6bd17a411bba7684e1f06be9bc3d1aca25d50c8bd -1d75597194cf87c9ffe04ff28bea91b5b9521fd356ed9e036466137586ee -f0a8795486438d0d9707cb2854f12963929edac394c562235ca71376d938 -e4e1518668180b857d75318bc22e9f0683749047e7649f9e20b35204b6ee -60c0d47bebf53179a083f0b4cad5b3327a3faf2cf03753e3e46c05773629 -7e9bb305f603369cbb568350b2b5c6d23a35c551e0ab28b082e321ef4ed0 -e2704d35c75b4750af782160c2f2e9aab0e14e541e95b64ebedd66db2c12 -a8935a60177cab634e20a8871a3a72f4b21c3a34d9dac37176a321c2ce3e -e828d140c8445117e7fe4738000c30ffae8e2a48bd618cc8813e38fa0f86 -92ca634d1e56010987483aa0f08980d91528df3d370ac724acb238e141ab -595dcb3da7a769de170edd5763078d1084e2ebefadf8a50a816b50722617 -c9539dbd68d9062b015639708dd900aecf4f15adb36339c05a9aec7403ed -771f9f28c60e52bda3ba6902e06334036c1dfd66d35ed00e3fc0bebf55da -416093b5cf512217c47f905ccc91fad879d63dd1380519a02025ddf15d70 -eaa1bd8cb6be67608fbc5c94796bd09ba35933f64c5e72a26db1ae40ef49 -af5e972fa44660588292b67ac670bf046cb1f5a7a0d73ffd6df862744786 -4a56393b0f1b4cfcfa362c74634713093161b29c94a2526b7138aa92fdde -b37a8c1f30a6b3837d9500b340515f0412e681f5bf36e7869fa157df18e5 -c79df3e6aca924d7b7dd2e0d5b87682d7ea6913b26397ac180fb75fabc1b -8e156ed542b9d8c83079bccd141c187f90d72694de4f6d08520d11cd454b -bd3c2e6d259694fda0c8decc724bdd650163b7f6ce1181590c06de4c0dd8 -536aba318cabf54782c919e07c2ffa1034143175d05deddfcd7dce6c86a9 -ec9bf6a4437da474aac2dbce2c91aedc20043f179d5c9120f3dfb1cf6906 -c27f2ec68cd75035c283e1672ea90d953a23a1515c420b81c3270fa06573 -4d003eca1bb71a2dacdab67e44f47c266c2ea1776648b62bc110671e6eca -4546d3c72c8acd956e10452c32532ed51bf3d0518467fa829efd9c896e8e -1e5c7ff6da0b51e872e403470affc95f25e1d2b9b59ddb0472705e14fdc8 -fc2af16527188508be10d098372cd7eb7d62a85c8d8dd1d0f55ae3ccd0a6 -5dd6bf776dc187bf4de409d5db3fcc5a6d852848a251f4fb4e01dac5e9b9 -587fa8c46ce03689709008b34dfb3dc105def80a1b515abcbe06e73fdf7e -7136e40cc922fe9a9da1726747e84427f288d934747b6c587490734906b8 -a91144ac82a57957cffab561714e1ff5148a39499dfc8cc96bf5d87ced17 -825e8f80cd943d9a73945fb8bc51cf1f9cb39c605491c1bb8f1c4139974a -59471ead310d041b1ca1ecd5e9f92007cd8243cb3fb1ec5256444699a9fc -ed6cb31eaf0912c16fa480a1cb4a8f4a9cb6a4d9a9903d1e2f674286032b -489b8a23ac4719fe435a9fa2d79abdbaba740e69d5ed611421b1aefcd06a -362ddbb7b79aac41e3e90657afc0b87a6e8c57ceef70a628efe19f568634 -50f47b5c6d95870039caa3d07a54e58df064bb5f59dbe9b9a2c7c84d7e0f -32386309560a0efa2cbfa27f861b208b2df4a062ffe2c59c057296aaf5c2 -0f48ffc9ff0692f8cfbd6fc6ed1f3a14537ba40d7267e6b5f69c997a949b -26577a9a99db3f53167355c4967dabd522292ddaca3c537bcf303ce76add -eb99f6664227a94d6a698dd5a5d40008349376067d057e28e55972264502 -e035b1f5e33d7b3aeae016f9be50f2aa09aa138d15d7af3c1ccb805f2d5b -cd4e9b2b5c288b2af4a25abf0a9093749377c9e8232ba1af17962f85064a -23b0a13f11acbb471cc700f9f1b588f72cb63d3d1a95a93502ef74ed212a -c452f1a84619bbdf61a1dc79c0d9ba29c7f19b400f682cf66f7705849314 -f5c8bbf973f2c53bdb060932156bf2c9cd8d36cf6271075500b0e3e6ad49 -958af46a9dc950f4c29f1ab5dc0a85924f7ffef259f778459c80118b1eb1 -ed29208d1145b21b19d62f755de4972c57a09b3decb0a8096ab025fe6b9d -be49ae35394f0ea40d3693980f97f712b27f0e28d8a549acbf1da63518d0 -374941effacf63ac3de0523cfac0dcaeb690de5836741fe58917c7ecffc1 -95e7b560a3e763aa70fc883751bd60ea0a0f893d8e9fe75a66c67e202c24 -84f66708ae74413c0101fe0b5003be20881345d917203b582a247e6c74a8 -1d0479f317aba7b9dbbc0a92e91c51fbe8775a44c57699acc9da84ad60fb -9629929d1edabbd70b4ef9887ce4ec2469f154fada42de54240cf3302364 -7c492ba17e6936a4d85e0751df0945463368a803fb40d8ded22abe118250 -86cfff1878abe5b100bc08b991cda6fdfd579332360f0c3374842edce6ed -e43649d6702f34668a29bf387e647f96d78f33395e8d4b3521cb4fb0956d -12c924c16eee798cde68e319a358cc3524c753177d976d4e14a2e0cb72a4 -80cd87bfb842060b1266568af298bbec58a717c577be73ad808e004348f1 -6aead32a3d57457376ab57197534d6e469ed24474a83618f3ce21df515a1 -22918f4b62c642de0c8a62315ebe02bcfc529c5b8f7c127085c2d819e29a -f44be20fa077ee01a8d427bbe3d97a9d2bafd77f17835279bf135900aee5 -9bc49582b18d468bf93e47ce0bdd627775264ebe9e4172839a444f928580 -8c95895b7e23592b2dcd41ee82e966c26aa2143e3057161511796e980998 -1f2e4ef5868b3bf4576e3546e6407e35cdf14654bcefa7557d09407545a2 -38173080b4771ea52054736677a8d9749a2b22b46b24fbff93c55aa2274b -8c7ddbd751bcaf1df00ccbe1f24a80622aff192fd6db2238db941ec44ae0 -dd73f6b2f80d89bd0aa30c038583deba14913d38a7b61b54522755e251b2 -aeca62033a39ec1143b2b960f9cb87f748428bec3243b8164f07d5ff72eb -f2ef69347bb933241c2401a96ba5ffa3f9ad060c41f4e6bf7280af65293a -bbae49d723dbc4be61d7e13f7a5931a697e7f2c6582dff416341ccf5a24e -9a53686a1e13bbe0bb480c19a4e72a5e477bd29f39dce1a17f63f1e8c696 -d5f8855cefdbf7ce681c7d6ac46798ca9bbdc01f9ad78ce26011ee4b0a55 -786bb41995e509058610650d4858836fcedfe72b42e1d8ba4d607e7ddbbe -3b0222919c85de3cd428fed182f37f0d38e254378c56358e258f8e336126 -9b1f1acd7f387686e8022326a6bbc1511ed3684e2d2fc9b4e53e83e127e7 -84da13550e593bbad1c87493f27b60240852e7fa24392fbf3f478f411047 -3f00a8fdb6dcb8aae629dc7f055d85341d119f7f6951ae612ffa7df82111 -d1ca48306a57a922cf4c3106f0b5e87efba6815f6de4294c7a0394087067 -677889d22a3fd86b0796200300d2716445078027fe0c0b05c86ac80d2095 -ae874324ee6ea3553bcb92fc1522a6d1524f6fa22b71598fbce784a10b5b -61e50307ef4409ffb7b38f27800f2185140ed08fc4ab396050b068025a9d -e4bddcad201e72ed9b41c4ffd4cee743c9c2345b95c5071442defc8ba5fa -9c63c56e209df41d10d93135a8080f7cccacf67e0b0ddb3e0a31df32b83f -290b3c536e9949973cdc80aa5c8a4feee20290a95f68e59f54050192de42 -f27464ee374e4d2451ee8708933b970402c90ca3070843a449d7c3146347 -1efa666a60fd5cbf55a47e4a3c5c318fc1af944d58d32690a2c7eeef09b2 -d94721896e1e3e76e44a8efd524ed5d6f5eb9da093d277441546c6828745 -ad71b6c13f653dd631bc6fc55d0eb4648b7bd9c0eddb13222542f2b6e8d8 -b80bfab4365f4199a41ac690979285d917de79359a183e6fc254b63e6408 -6d33e3c029f472f40742a99f92999f302f79994ffd615f1a848194cb56c7 -12146850f5e400303bf5bcd4e5fdccd1fe2edf5352d525cb15d8327f45a2 -6e3ac276dc8780c65724d28dc6bf9c7c985840070c35e32859168890d599 -a884dc2a90194cc2e9cc6a20c6c0ee11b20adf3aff01db48eb8dba7b0c81 -7fc10cf5a66e8171a2823a4cd22f0e80c82011ae56dd895ae2d3ebe84ff3 -d521c31453e0909cb9b1cf0b030eb6b7059ec38038cae12d0e1cc4b5b3bf -e6c821faac9b8792441e2612aa1ee9318b71f9966d7d3a64abe349be68b1 -744de7b212f6be73a0e1eb2fa30850acc3d9562f989cb2d4fbfbcd5d3ef7 -ba55717da1cabf197b06ee4d8650e968518b6103fbe68fcd5aab70bdd21d -66f09f96208db67c1b345672486657295a39a7fd689b2c9216c6b46a29dd -1283bdba295dfa839a45b86c14f553ff903a6f7a962f035ce90c241f7cde -13bab01d8b94d89abdf5288288a5b32879f0532148c188d42233613b7a1a -7f68e98e63b44af842b924167da2ab0cab8c470a1696a92a19e190a8e84b -1d307b824506e72e68377107166c9c6b6dc0eed258e71e2c6c7d3e63d921 -39690865d3f347c95070cd9691a025825421be84bd571802c85e2c83ba53 -841223435a9ced5dead103b470a4c6ae9efcc8b53331c61d0e1e6d3246cd -aa1b0da347685121196a07e97d21b10ad34e7031d95c1bafa37b4141bf33 -a6be401129dcd64086885f4b5f1b25bce75a4cc8be60af35479509e64044 -d49c8a0c286e4158a5f346ef5fe93a6d4b0a9372233c7434a7a6f9e7ea21 -30c0b4b9f62e3a74cc5d2916ebdaa51a1ef81fceb6cf221e70002a8a3106 -bfbccc2d1809dde18e9607fcaac008fabb72e8c50244507f4013c5a268a3 -6135ead9cc25362c37aa9511589f18d812e6039490f9c599f44e88754ac1 -4f6c1841d570efde27958c7f1b2c68772584e1d12fea252e3a6ec3b051a7 -6faebbf6f5101978e24a9ca927c02065e8e49150a55c64dd30757e8a33d5 -2a788437a9181efb47414dbc22fdeda203d4122137bd045611f68314e12d -1d6a5ec270c8919562c03e3af7b0e0deceeddbdaf3eab8fb5632e44dc1e8 -d46e2396b0236a46659164e33709415e7b347f7f7b87a9224a189ddf5178 -2cf66c9d385470a51efc88696176f6d3ac3b7b95fa074c981194e22981f5 -1d925f980393b7102f1f836b12855149ef1a20d2949371ddba037b53a389 -7617c257bbdfcd74bc51c2b40f8addfe1b5f8bc45aa4d953c0d1d5f4091c -6af796af6513c820499969593bfd22f8c6dcde1d2ee2c0ceebb5bd6a1ce4 -5fa61094e932b380cee381f4485e39b4b1797f2a7d8d90bcbf89b9cb1006 -2d50fff083743bf318157caac1c0179c87c03a2857fc002979e7cc97feda -966b09ceb761d3f55cf07637256c6aa8b8e5cb6aa9739452a330afbe7082 -975ee39fad5e8106e8ee05771157e92d99003533d922ccc37add065b6236 -7613d039741f99edc77c230fe8d1baba720a185186662376b947bbe1a686 -4b42c61ebe1abd40d890751ab8945c629de3b6d2a49809dc693f9e397097 -cf1e568c258081242460af2de0ca44b7ba2734573967b3bdec0e5e64598c -cbf41e630d821491504f414d9b54a3100dd5105a141cf61bd3ec41b67368 -c8cd366c543754ee800ffee3d19c9cd0d408cc772da10e4d8134964b0a61 -232e2dfbeacd0fdee12792504bb327a2e1fc44127f8577ca51d380a760b3 -740e6be46455cbf3917b90f0dfeadaa25d5d9f66cda43ebf9f75e0191a06 -25ba29666bbe8678822a453d4e876bad4a6b0d4b6cf98feb60339c9eba2a -dce4ef7faba428422c503d0210dcf8d884ca9f5094aab9f3b1a2238b569f -444748902907cb0d9d7ca33fccdd0cd29bc68e44f7bca5092be6272bc949 -baae5af92c302bb21f91b6ea8463265680f7c16f45d8ff35392a10eab87e -296f3af4478032b5b021db8510deb617941130d45c46fb3647d94b162fe2 -2738766fb6d76a06ab6803818b27c5ff4205ba668f95b5ec5ce4ce6da545 -c13ff56f417a4e0b3b8554a1e2a985a167e168adc8c4db28a601a80ab451 -91bf32acfd8d25c39c2f17fb3bca1296d3d160f25b43b4d6b94f20ffe012 -b779339b12860dfc897b366e3d400e756f4f9f4d2c86fb9d94c11ebd1450 -eaf720056e2c39529331bdcb104d113b42c94af2c6a5035750b7ae7fdcba -b6116d74bc07a11d4357ecf73d99221dad5cba4a7136425c2a3ac0e092fd -606a4ab722195e3b7fdfb5a5e3ccbb85fc701c42bec43b54e964dff3fa04 -193043eead7681cedae9cce6919949ea60ef5630c4b9263c8f98b4bc74a1 -63ccf3d0a0bc1deff39b800ac90bd734dda7ecdc73169ad77e129887db80 -7a253f8807a422eda8a16c9ee9bb8fc0942634bfe035dac9f7e36d09844e -39477c043399db4d07b3617da9d6eee76d0fde9201da98b906050748b68d -8c944ace3c96e90a3c2b63eae27b9152cb7274fa336866d71b65a57f1bc2 -bb1f482a67f3993dcb3ff24abb0223f9a026c81b2b33127a1dad8929dec7 -5d46bdd790eb1addd771c5c3965a2f514d3a128117a44560cc10a729bade -4e6c86de7c09a39602235c803902e34f5c176b18e127d71a011dd9a3a61e -ebfaa4a4e2a5651be6f4067e5e09bb4f3514d67c2129e4d3ea9568661138 -1e45af07bd84f883c70577a986416747f3bd8d1bf86d3d7b07e8a350899d -3c2dae237bd5ece45faba7a0ba30fcda7b7eec9fbeaa5a94620686d1e403 -1cd2512e8d89451c7bd8eb432c8862023d66f3f9fcec0d47598e2df59525 -d673a5ff493d458748cd6341f161a0a3e8996ca5b496508578fe4f653924 -2ae28bf4b7397c02b726fd5f9d8b898938bb668a546be6e42865f4f030d9 -5faa289eb24f7b8e249b224a95a2245605d67417a489626df7417855b8d3 -1c0043cadd2b461d32e1b39ccf409757c37b68f84e752bde6b5bbb847bf1 -57ea3434802def983d6ce5ceb3e9fbc4911b5484e99bb94dc3f383e50672 -0e85a91ed378e352838cf02921ee0ea94be01b5a60f9b1f58fcc1b4f527e -43725de9b9dadc3ef462fa279bd7138095d4cff2a0563039f71e383430dc -f628dc9611b2e3db08fb2da1d5383dc1a3c784e1e64541fde1d9d7f42505 -de96d3d0a401099fc2879af0293b0eeb143b78cc221f670c0479bc150047 -0cacb9a282e334e428b527acdfbfc56e6aec8d4d60745c1dc000011b6248 -d9ab4a17dca7cc74e17d33c0641710b02cb1edb0addc6be214b17e9f845b -2d9c8bf03c19e131e00f91f2a393b5f2ae7c3d4ae9021c4d7891d84d5067 -377ce92836e42eacd7e540824f7ac95360ce116d41d17a50748748971c82 -27f089a22ee0d21940de854f737547b73c7517addd9bdaab425a6c2908f6 -87dd990d6cba4d84308bdd4c4435a6480ecfa1a14daabd4d8e2398178e48 -de28b84f7ce4b61d2e6e64fe043c29a941f6de7621ee6f6d8b506221df05 -db238b8fe4323cb5f259d4d3d9c94d4ae1ca37d6c34345489c0284171346 -e9830e2e3c6c167238a7ffe0989d3eac870cd44102cae139469b9d909b5a -9c34792f693ac94ecd35d2277080e30a2d24b50391b6f2a3d3b6c81f7ed1 -a7b218903e7fed7a63269e27d793a2e0b40320ebf447c71f36d40dee002d -7257f43c8add31edf2c571123e46fdb413e007cc89e99b6f98d77ab38bff -cf140f787e45ffb2c7cc4ddbb59a4e32dfc36e2875f204ac851d757c1236 -12deb31324ea4c201d27fdab46e9f3988ad2bcfb8e9cfa8c487831a9b0c6 -60b20fb66b4c77f52359ac96f3b3d189aa0571c1c53db06ddb10f08882db -0b1e93e9478d4c75626c5fbdbc6044c4d82684b310ab2af144d12bf36f1a -c0bf6249d1da9ab319453594cb19d0e93c4e047fb49229c0cce76d0cece4 -2e76fabd2425382afe707db032cf617b046a59a2fc1bb3838d98fd5c8053 -ecb918bc14762e4ca45027623988f434ff4cb08bc9bff5d7de21940e3e03 -1ee042d9c30662aa76f96213fb5a92047af60f320e4660eadd1ec19d0086 -072f2202af5f219725f81882f10d1e065a8035a9946d0ca0e48a5e7dcf61 -0283b834eda01e7d94b3453830daade2aa6c947989b290c02ade0d7b2620 -813ad177ed82813b6a985d5c0a2d42419bda763d409da085936e33c817ae -68e5467eddc30be172de855a0f7f5c527555b3f4d942401b450f08273b1e -c5b5352fdb8562a71f276284cf7c27537e628f94bcbffe8d669ea2645752 -60830f1e65e83a2204cec393f6d92d4f61f317471b4b93039d298ca2cc94 -eeada0140823a2bcd1573e732e7b4bde7368f2ecca5961ad547f554ae989 -98d87b7e5d07a85c382bcea1693a697224f41eb8b406bc6a0c3eddfe8b5c -f25b11c3e4bd91ea7d6274cd6b3ee7b8f18cc3fd502a324c645568dce9e0 -d43caa61f7306fd5488fcfc439d85f8160ebf0ac90fc541f9c74d35d7833 -09309807a639477bb038200738342e50136dc64baa7cc1b879c61f7e1b90 -e1f2bd4f6e54c4dc97b8e4adeb102979203a31fe26a7f58c609915a95abc -4acc263179423f8ab16b04272d5592fc536f29a45cbcdbe15890f119ca9f -c7a52eef41dfa5c4fed087eef8e698ba738e300bd58f2a1a10da1198c1f9 -b60e2032f8384a86aa84027df21cb87977528e3bb9bea1e3a6879c56402e -a29063afc6ac0194f4944433f9a5872cf0a2a741382d7f3c0ca7817d5d7c -4b8bf53af0f18b1eb54480519cebb61d983157e039b13025e7980eb36f54 -3451bbb84e470ffd0f98eba80c74f238729dd6278294388a2e06de68a719 -47b6d478c85f124d14aaa835620e49b7f5a4f21347302c0f0864f7ebaeec -d0831c36187cbe9c848736764a31056d2cef27c07cca00033dcddca9a2f3 -b9ebf28e67257b69cd38bc23c711b6a2f6e4dda9bf5a19da275e6a8d683c -723bfbb95a90a344a6f421f0b67ae84c74652288b0597e4c86c28f73808a -77455f2948e8df634c2d14f221626b019033f9230c9167982cca9ae6dc37 -aecbcb49fd9fc1dbf2d11bba7187888721bc42a7f47c23e07d2fc5a7a91c -0dfe255a7f9d17e69af1618502a6b90b1dd748c7eaca1e1ebe8b861b04ff -e5f628f47eb4e7e65311037d7a5713d7cc3552dc85f452ba74c4f12aecd0 -d72892c940c3325640d62fe3bbbc71361dce6d54766e1fb99dedcb2d19d2 -fa6fa21f9116e03952ebbef659816a62db51a9b5b3916ff818518774ccd6 -79d44100d7236f211f36fa80a4cbafb3db76ba1e7e7f12082b0140eed2cb -5e793e24501715c6c170ad4f856a4bf16bb10210025156e635264d3cf18b -1fc1e8cd2fcfdc2ab1a24af9087975bfcf6fb703fb36e288e58d0d2ffc98 -bb4318001d931ad6161dcdf8984e6690e0f6bb07af81bf07445f8f57b355 -6b960d24e7cd152708489e4d953ab6a155a757e002ead97585e6c5333d7e -5aaab2731f047f3490432e0ebf3d0d628eefa8c1f665b9c86aabb0706639 -5bc372e16378f0d9b439c98e7bf87be73e934995d58e4e70d3ae9a5b54c8 -87a19f2826a772c39d41805c642354d9bec75b065f148f7c1e435dabbeaf -e4a5744e3f2894a928121ab069bffa3218a106a9dbb83971353a7c7a5616 -d9da66fbb908173f9b07aadcbd4d112cc353e7b70476046ce5a92e86eaff -4eec40acc840005f51f55c9f5874216851e9cf3fa431d95d3032e779e356 -4bdce33966a3a798b170a06c4cc9f73700224c858c36bbf2d0326c337ce9 -46f69c19a84187fa50afc5b36010f9a7612e3a25e846d49bb907af9505e7 -d8c78748d7dcb501bbb3d6603e829deee3784f2f3ca583d3738d6d2ecfb8 -eaa887103606211a3c1b5cd74a3e0e96fb57da91baebaecd3669661e7b1d -579ba41928a40a7028acff6cd409e601d23ff66ff2c8acb12e535360d727 -60d2e988d801930e0e9443d60dcb9f378fa75d58d73e6a3b6e5b26407c82 -67d50ad97787f8a9b91765e41552283cb67e43e59bf71cf08b9755c8ce47 -0cf374832c72d1e9702b55bcfc8b5a4e966d5072fb2a72a2108574c58601 -03082ac8c4bba3e7eeb34d6b13181365a0fbd4e0aa25ffded22008d76f67 -d44c3e29741961dbe7cbaae1622a9d2c8bca23056d2a609581d5b5e3d697 -08d7e369b48b08fa69660e0ce3157c24f8d6e59bf2f564ce495d0fca4741 -c3a58ec9f924986399480ee547ad1853288e994940bd1d0a2d2519797bf2 -8f345e1bb9cbf6997dae764e69c64534e7f9dd98f86b5710ff8b500e1c4d -f509da50c64e213ebdf91978553a5d90908eb554f09b8fc2748c9c405903 -e7bfbf0ea7e84254fb6735f09bf865244238e5fed85336c995bc3a3b9948 -947a6eb95db4cd1b64c0fccf82d247a2202e9e7eef5a550557625a0192bc -8bcc9e461e52833f6b8729ccd957d5c4b6e07016e864fc02b792c7400ace -d0a8f43c755f87bba6e5c6e1022416e5454cb34a19865d951f7aea527760 -53658cbf306ead832244f3062c39a0a121a1157a8e47008163c5bfc88197 -be16e9a1ba26a035a16dd38cc28dffb666dd4ba7356c66b7bced9e26e905 -4ce25f6d36607d8f5dda1e21ac96a815bb2989f01130ba1aca9aade554fe -effdfef5d6b0d2a01aad92f599f6a12e121010ae6acc6f150f19e7305271 -97da761b07530ca19b84b119e5edca1fad18462143b8913d6b3f6864b713 -7a93bb9e1bc29c09d660704e8d8292c61072ebfe35c354a2342b2458a353 -31d043874380d439388e46688a53bcfe01bc190ef1a6b5dec9d40aafe822 -261b28bf3e2d76f3dc4302506ce3387b4aa2a51cd4ba1faa2ed1fd7df664 -6772fe9f83d253451eeb0448b444b8ca80cc7cb653c2d1eaa0de6f2b1c72 -47e6d24ae72e620e200aff83a557a1aa7a0ce0a9cfbbeae03c31d8cbf1d8 -20b53b688ed2ffbd83418d743ee31e3d62216ac7be6c12bc1917548cf670 -d69fd2e78d9f7786ada0ea30a6f6d9fbd1f1406337151ffa1d3d40afbe03 -728fd1aa2fa8a4f075796b9de9586b71218b4356fb52daa01d3c18cb75ae -d4d33fc809dcb6e3dcf7aee408a0cef21353d76ed480bf522fdfe86e0e0a -b7d097defcb793057f0ce98ea4989a9b6787b14029a4bf10315a2557149a -fe9c91e7d825f7518b343fb556f0177a8f6ca08fbda9913d52997511590e -b9942c9813b4cf4d4aae4919401f2fc11fef0620eb5c40532cdb22d5fad6 -919a3a710de6c40d54993b5386636499c866938e33bc703a99c73adc228d -95cac73ff4f4a275c04d0d787b62c6a184dacc4024d23f593e7721be232e -9882fb738160e52ab905f0ce2c76ae6ff2c8bbe118a1acdb3b464178cf01 -94bc6a50df1090e9221be11e49f254b06c3236a31569b947ad041d1c6b55 -bfdec3c18c791ace0fe2a59504eef64a4eec4b5c8dd38b092745e0d5ad29 -276bf02c419c546627672a5764a4904635bff86fd0781d36fbdf13485229 -71f355de2b0ad250052f50ad70f61afc870ac7a816561d3232b73360d4ab -2727b2fd045f254c782bb3f1f49d94c6d625047071b7e32da5c6d21a86de -9283fd632074430772bfbd85e0c9ccab1dec16bbc049c3e223bec1b65c8a -9e98cf58b30a74f74f1a842dc91e30c023498e280ac55edd58f4cc731d81 -e443d9b9efdf5fea63c9f357320e01b8740eedaeef2495cd02eb2f338b3e -674fb074cc497d7b1937b188da857c2c230e9a931cbc00c85a7a36fa80b4 -56588e1bbabbe4ef429a6aef9bd4eb89c5752421bd049aa13f4dcf9b51ce -2503e90bc118fac78a25d187353d6f5d496cd6130b337666f49619cea985 -dfbeb7e49c67c1e0f0f8e9ec8ba14624ed0982dcbb69415e4b3c8ddba140 -397eb1fc1ddd36c94c374f018873ba41109e45afa51f0e691157d5958c06 -26fbc0903ae25e47ee372389cf65472a3e4d9769550bdc42c0b72f9a297c -d5d3c16ec67e06036e740ab664abc9f10b9499269b73ad3678daf4474329 -c2c7252c1f0df1e3b5e8f198dfef8325cb1e7e8057897a3d7fb5bb5858e0 -cfc0c115bbd7362d8e8ee41862af6eeda681cabbb06f72ebd2ae0b0be45b -a9e1be83f1da30687a655e5d148fcc17d9f53b760810a565f6d2f4cd5da3 -5434116edef756adb4d3df544a1de593be988f2bb8d36c34deaac7d9dc15 -cba49764f1e03aa09fe21fcd7c74e3d6487ebe219569e019f10dd163046b -c1a3cb2bcbaa8558197cb2c18709a998b4efa8ab8c9a71d2ccf942c17662 -1b88dee6b424165d6ce10ac48375e760983818e0085276b1674dd41042e1 -a01a8de111c903f74834199b3230bd475d92c6226ef74eb1daaec3475a6a -fcb47644a17c7e390ee3b16bef1c1ca6c55eddc44fbefbdde525921b3047 -0d76817bd8ac724739a8e743eb09cf78e88adad527d4f115b8a32ed4898f -45bab3eb802b8168aec061e3ecdb026c056fb9efe7e2df48bd516ccb12ce -00de08ed8be4ee0c41f40f4c8f64483e0ade90a78d6d4fe9203fe0b97c60 -3b2f8882bc15a212453c691c52d00fae8a3a26934ff8acf68d4352eef75a -0b10d938e55b7333dda2db0296a69e9775bf82b1aa6d684fd9080fc1c11f -ab4369c7a95a9504063db900a6e345bf6dd99be041230b2e60cc86b8c345 -1d84a9c2cb4ab6d74d63dd43dc26eb6b384f5222796d4083dcc3e1651548 -d9469f09a33b213a33ac52a6a2e23802d8f8a75c01a607940daab0051410 -73a88130bc192f303616adb113c0051b65e12086cb319c0a5323fa7def40 -402f5f87a3b2c2cf0e92789985f6775ac2743e1ffe2d0668291059740d45 -43bae7a2897e5e658592bf5a72966097742e0702deecb0cb12499eab701d -34ba37a08346217a415e44297a181bbf3744f0a49230ad6f030e11462be9 -afc2ae14e0587bc02311b48b8e2122c28cdf14414f3680fa52dbbb63b17f -6ebe4a1204f3c5d6150cbf89a8023890383153838d4dde77d4c8b1b78823 -8918c564d3babfe58eeb154307dd1997f5ab7105426e35c279008b2677e4 -695c60f956b348799c04b734338018fc27f7de7ad9d73468fdbc5283bd14 -c066ddad9a3562f16baae15d72d7bfcb409e1c874e9db1a8cde233b282b9 -6e76e9c08d85ddfbd3cce7e64104d0b0e95291bd91f405ff82f41601ee20 -8471e613fbbee67f269e4e954c36d1d18ca9880b7cc2b08fc990978efdc5 -1d157deefedaa765c1e26ee125d4a2514a41a3b95e9151a824532d7d6486 -35ad622718fe71219a697e94c2e64f26424cbb767acdef5cda70e179cd29 -b7e318d1c6d3ad26fd5fdcbf2fc221301cc1f10f5ed86b40a1a6bcc01c90 -eafd65183e75609610637b99fea57885efe76437df02a2ffc21223d039b5 -74955d9a54ff41980eddaa8768c5ad883a0c9150877392b990d63c6805db -7b8d6ab1358cbedaedb6feadb0ee4fb8f9c1ca03a3e755a74227a8930bb7 -2ea0a00b48fc626fa14d7d48624aedc31c556f44e982f3ccbde7ee735f73 -629ab1b65bcbcf0a3586a920477e8c960219802fcb1bc3a179032b324f8d -c424899b38275886cb5bc771f26a0880767d49cc23426a40a4b6ff8fe48f -d747565fc537565f6d7fd08706accc60f5fbcb45bc785f45ee9b0812366f -ae71b23ec43f3549c8224d78baf18719f05108d5741e681457ead8abc050 -462481771a8dc6cfeb98956e163981a98c59ab44d90e9c3a946c453b5071 -db0c769f7fb5144c7ab0c9ef1a6db1addcde1d4ae1daee1b4035af256a04 -df53926c7a2dcdb94caaf12f986e20929ba4e396f3aa7c93a7abaef1294f -5f13a0dd3c3aaa8fb38da3e15daa32163b7437af683b4f5e64cb14aebbde -8c69ed2e8cdbfb213fc8129af29ca2c06c8f85a5038d688d1fa5d1b54ebe -4dea81a49ce24131f8e6702e7aa4e2cba078d5dd373f894ccb275f49c690 -1dc772e1d2f5fb3fe15dbfffac62c87110162074eb72ae4e5e446bf7e650 -a554178d0d64d3c07f330f0d99e99f2239cb1597f2e5f443854cdb0f5fab -b28fe62f22e7f3419d017980f325351bb04f8f3c3dc57fee03cc029bd29b -202308d5a800ed2d500d41ace8e54e2557bf25b627883beb8118d800eb94 -f4253f855168f7fc8a2d29c5fcb76bb90a6c4e345722b8991a854047f46e -4e97336be85470b6be2b9ba573dbc4967ddcdbfc3b6fc35b0c7f3f2f570c -55dc3fee6d80bc6f46cc7e4d86a0b86f6fa61d062e213d9e442db63fbf11 -d03165b44572096995ed342893bb672f6bb55ff8fed944667995f0f89a48 -a904c47420f32afd14129c6e2bedffce1f07ea69d550b6909bb5beb4aa08 -b0b44f35e018ba5206fdb4df0228462c1fdbb95a429e53eb27bb1b0490db -f07202c3608d0f4ce08570e3d6aa3d4581c569b57bd8c1ea0e4ed3fc5497 -e316ecec06e6be582d9170d426f6d22d8c7287b8219945c124941ca8812b -e97efd9105eb6999edc0665016633b3b48820df736125b7c76c9f3a67d93 -8a2a0a6b743fd42aebc46a0249be459f16811ac9eba7b63bad7c2e88f175 -0eff8da5faaab5659824f9d19b3225aad2ac17c52c523414d3031d08a926 -30abf474fe02a32b44d3b7d9fe0c19aec16ca6d018b71d9d395ffaea0788 -0d4501d7cdf0f7077a2d63303d09083080d67f1f714a1b271dab9fc9866e -4b0571a171eec8a4e351ba2d02438cd108a33b1106acaad0ccdb051061ea -7f40543748115f29debfb4be4b42cae8762d62114ec6f8ef68c478a8e05d -ecfa18b0368428efec9eafb2353f95e3d71e1636b9d9f94a77e692843255 -698576dce13b2b858d2d15ee47cdba3ed08d64b77ab46dd29bba6aac2106 -ab847de378cccdaf35c64e50840248915f4fc110992c493cb1b9cd0b483f -0f1abf5e9b018210b477fea28234ffbe5e0bbe01338e0842a89f1e00a0ca -7cdde0b2d7c324d5e17d8d3415ccad703507497ac95360ce660b656e5f66 -72a2f50761f3d02ccdc1d5692d7797699b8e2147cfd4817c81a432ff6a5f -39cc54927fa146cbed56a55f85f123c0a94b7553a8819b329d9dd122c502 -94e3f6314d5117db89ae7597c4691b6c542979a1ca3d26a8e23d3eb698c7 -1841651e08ec771cfb974d6613f2143872c739b62796bd0a45172530793c -28d93a65b59f79c245248d2c09428657a35b0c0e367bf7a4a4f0425b3f4b -485d9f402e164328a4b963f456829a39035c00283d2e4fcb71a42da6d42a -d46cb751287de34e6519c60bb3f1a6ba91f7bfa21dca96ee712af5681701 -18ece8a0535d9ba1dd4bd835e004a2f38c5ba43c9b30d17045e5649fbbac -188922e442182d4bdafaefb39e00106a5a7765f3d67850471e3629e526af -8691f935b57bd38465665204a214fef1006ea37dc0781073ced5fc042781 -93650393c3cadfddedcc5550ed483bb6355f54600e9758e647f9c9711f1b -e7df05d0e50a698615307c18f6d4886f50188011ba499d03831185915f3f -77c4b9ce708d78423b110776aaaf90396be0381616d1e9b0c1dcf68b6396 -82399da2a7323bf42ae5347599ef4ae9e5c135522c5ecb87e201853eb899 -db60d24acad17d6b7c2c7ea4dc221f3cb6d6caacd1ac0822ea3242ad9b4d -d15116c3874e3012fad26074a23b3cc7e25d67ef349811dbc6b87b53377f -0cf972040a037ecb91e3406a9bac68c9cab9be9a6bb28e93e3275b177cd5 -0b66935cbe8dd3d6a8365625db936b2cfc87d4d6e7322df3dbe6ccda2421 -a5e5372566f626a5e9d8bc66959e443286f8eb4bcfdeb6c49a799f1efa69 -63260d0ea2d51260baba9207fb246da927fc4c89e9c4dd5848fd4ef6f81a -cd836f5f06ff0fe135cafd7ab512af55a57727dd05a5fe1f7c3c7bbe8ea7 -e6680fcb3bbbee1cf2e2c0bba20185f00e2dc3afd42f22de472cdb3eaa5a -ddf8c6fb3682eea5548c51ddca25ca615221127b4438ea535ab3089c9ed9 -b971f35245cf831d9461a5da9d57bc4e5606d26535a7414cef6aee2a7b95 -bf2276044818ee0f3b0a16532934b8b745d8137b42ec2b28fae7d55fc02c -9ccfa4e0055f8a4be96e1e235c01b8b6ad509b832a3e90161e0a449934e7 -4be973c939b31cbc19dad4c58e9be89d242f0ce200548cdd4fa2081ab3f8 -e01f358d5db24b7a50eb2096d833378921f561f132cd7988708ee10cffb6 -2256201801c667e176b1dfaecde9756d725bef093457805e16f550e8a7de -87ecd46e5b09646b73ee74f890a36867636911e4cda2c46a40e7d57cf297 -9696046614c85b1a47ba55c60544ebd3ad7d750d003bda56dd7eed8c4702 -f8b319aaeef9d3cdc59b3e63ee93c6e1e857af273eb90909ecf36ef4c276 -895c78aa762e5376c5c542f854fba864ebce56e4b0207091139f053c2c08 -3b7ddcd0a9909b52100002bc3f8c47bcb19e7a9cb58b1ac03fee95e81195 -072d3aa7c8079632725f63425a3550a947834d29ac9a26d0774e90248e18 -996731fd9aa53ab62b40ce557d98e874b763d9d629a173f0c7babfc00ae7 -82daef5f00cf3608ebeef403dbbc19e16a1d160b889f4a10359d9eacc19d -7b5f126b31720dce7fc35ec861dfa56ea23fa18423ff4e8fe6e53fc6ba16 -b95a2b5dec00f614e4f835281ee0b4bf549e7e882689e0b445dd46fc40c9 -090e5575fa2c34b02a51ad0bccf6a7bb83ca3b929285e5e9fd054b72c47b -733a66c5abda526b18b2e49d0746e067e63b948a45eab2f4221c5b62ae21 -a5d9d7cd8aa9eeb49588891d22c56b14b55ceb6488f02b73ab3b7f6c5555 -b75452594658255e4cd58ac4815f2e1bc3888c6777f62aac2f0a57d416c3 -765c991f0f9a33d888aeb2d527b482c042ee23783a04a73ad13dfc590a52 -f3116f8296cacc7ab29b7d87e7864561a5d0a12bde2d36ee697064f41d1b -ca6ef2f801caab5295d19bf4c02b10c19f73b44635ba48a0806b967d7dfc -ce9a4850171a78532cb30020c0d66b3b1e7c75eaa7894904c181a022e8bc -9b2b8ef1202f3c7d36bcab4742d4a4761bb55b64da0d99685d319f5da8fa -132be6c0483f50e2657ae8af1e28f969440d6ed43eb00e95fd9e1cd490a4 -8646f6d008598751f7a41b43fbec7770fe591012b6b0c4ae18775ccc7db5 -de0ded2dd53e82c89648d46f0d0cc5d3ac5aa104239608d512a4353b9547 -04fe6eb7e73d718323cf9d748b8ec5da01ec9358267de12cc22b05ef0312 -e4b6ac5dbb6d06d7f2d911f20d527f504d62547aef136834b3695df8044c -383b6145e824d3931a602f081d9d656f84987a1ef121772f1f5b37a116bb -d2e77d4ccda01411545d24e15ce595db4cd62ee876b8754df0b85b44e011 -b82d76ce45795e6c2c58be8690b734a8880a074f303a70da4a1b086a6de6 -56c02cc7a4c25258eff18cb0fd868214bb46f972e26509f868d065b3cb14 -1c316898cf22293391bd7051ac3a6927aada952a8fd0658ce63357c07f34 -acbf8c99a5537da0023e901f0eb5547e1b466b7d982c8c539798b76ee2a2 -252437a81a37c3b63f625172d682eeed0b795860b2755f020ef52a138353 -003c61be2052cdd7d73b2cdcd26b127660a7b22fc51a6a2f6034f37e3e46 -c1d7f83f8b28c7c965993abba1d358362833580d9c63fa85d4cb949f97de -579fb6807b95a58b78f596db50055947dd0d0e597d9687083e9bc0266e86 -90b884b27f4094d8fb82ffdbaac4d580340a9ef8aa242be87e54b601af19 -87a48d267c04e371ae77163ebd0de3f5297b1060442ecdeac38334844e38 -0f294d4be73935fd8a38de7fba6d082c3d9156d7e88f2cfff0459377cbb6 -041f37a7e05010753b98e0b67d5827aa312129bb3c3bd883c12323756406 -d555720da8a0bb30edcfa760c01ecc2ba3b15fecccf5a10e9f358822e0ff -b64178fce2ea6a1105bfb72df0e4bc499b207ae26b8ea960de48e7ee7010 -b4e671dff795e4cdc5b43e81b1604d224f0616ae311f1208859c502c1a10 -940e7b9cd11be728bd3a0c8005ae23aea32c1b642812198a6f1aed32cb75 -97152b1340dd35ada1b81051e393d38f3740fa9523df6a83b8ca7dbceb33 -6e299b54cd998d4dfef804733c76156585e42b7284cbcc4047ba6b290efc -aa60953e98cd2b4bc2893857fa6a339f820142a52ccab0df09a2709df550 -f22e5921cbca408e7998cc1cccb8adf6d8f8b71e6685ae59d290fa33f5cd -664d73e434237424060f634262f04e9a71a977556e93b692ddc3aad26d92 -97dde71e4def64932151ad572af6e681082e9944ddbec6e7a8bdfd534233 -9ca3106ca1ccc80eab14f1655978b137fad8f399df7cbfa2d7d3d9675e0e -9afec37369a8ede2c93145ab3f42a375926946680c215fa16bf7416fc892 -bacd806cd424b9f85b47802c4336918f7486af2a03bf0d39b10169d35494 -419cb1ab7b8f407897f70c18303e91563b497d70b7181ede6aa0c3efe089 -ca6135b34dd1019b298e3677f8da61f864a67023c31eaa716c40cf3d397f -9a1209564c9ec759c37028079661d2a56374203c78b023ec61340bce5d96 -e477a4f77e5c0db7c0d1257b4bbbc6f889b17e6eaab045b8adef6f931e4d -0795583d60a6b7002cf61639c6f930671f3b8ac05a1c4e002f4bfc50d8b2 -3029fc4dce1b602cc3a5533336271bccc226559ffb127e3a562f92f89824 -552b9a70466d5a3c74ae515a222b109d490f26e8fc2d9d72bc8af6d1dcc7 -80463c7af81993bac2ce4aece9d95ab736b1dc73e32d1237bc8ec2b52513 -36dbabb4ecc7ceb5d18b02043281eb9a3bfdf19bc4853c9b1722ef1cdcf4 -fcec534923db2e2653dc48545a9850c0ac2e4594abc9f7d18a0bcf2fadfb -bf085d465a4d10528312f5d790eb9511ca01061c0d94136b99a043bcf278 -c18223b1e0f1cc062b32b79e28dec2dc59a0aaa4b5f3506923c83e6a87fa -08a1d941bb644c994491cf7f3b0e2ccf6c8a8ba89376f76dfdb592374f93 -528e78e31e0b18719346b9f1486f652638e3120687774030444674cb0778 -96385c41f6566819652d825dd58f9a4308ff79b45d7828dcbfebc406e40a -c46e866cb0e3e97d6ce7fcac19a9d0fe39bbde66c5f0cf775eb3b1e6d7e1 -1f67e7edb3d5c4facc85c916bf13322b56a0414ca27d145cb740fa2c37cd -8c142d9301f1ac3704cf6a8e93973a07fde5a331cf0cbb370c7ba555de61 -18a6cea0ecb2c0e37152390cc57e2e4fb3791ddbc383ee26b6f4006d0d68 -4880888011020f856a9de47f45440f127cf27ccaea7d40a3869d39ec7dec -ebc06382d294717644b6118354e15544fd4c6d88df9245c9a83b30e6ce09 -e2498dd1df488a019b179cb859889e6ad2838f749e3b038b280ebc8d5c3a -b03e8f15751214691edf0f86281e612d7ec0773c8a5d2b433266402df62f -fcc06879ca196aaf1fc73a5f01ac46b44d6cbe7743ae9a862c20445ae2be -1544f413d010280cc2941900bf3c42ec088cb21b44a915bb810e7666b545 -5324465c5943eedcef0c09128a995f431382e2062f5e39f4338c8eba1bca -e553cb60bb8f3e5038ac8073398c49f06dc734b18afa7921ea0d455e6e73 -db8ad9f77fb5ba6c28af6b4f18cbe46cf842c82d6c960be1520a5fd929df -ac7e00ede976fb2be0a07f659079a421fca693de89ce9b8fcb42b0176d9d -f3ddd58f921e13e216933d27b49d175b423751c451be7618eaab054d3b8c -23e8dd6fd60182d61e9b5c86b3b764a29a62f913ee7524d8cb33737d7224 -d95dc4bb8c2ad6397604a0ffecc8865adcb540e5da1cd769077838515118 -ebc9f0b988545c1881dd2e7a8fd73e11bd7ae9085fb4d45526b23a346b0f -e4281ee3d588106db5f7c386c488d8f2f4dd02d4c08e74c1034f987a44e5 -d39fd07538de57a42987ce290fb2f6557e8b5cbcaec168f5780927226415 -1e11e3667d33b36a793aa53e9e2d1102c9eb30cb3ba0ebac953e0227fe4a -3d3c0eb57e4390c3d35db0c41946e45be2830a1ae33fa25cf2c7c9cb4550 -ce9ff6c6e3d628fc7284daa6241604c90dde6339b7f7e7df3733416cdac8 -e5291357e4983d74d3582a490438a7fdb0af97001a31990b1de68e6adb48 -917daa387e647f9f13312db57310c7dedc2a2ea80800b4f4bbaa99c6b7b2 -7ac8345cb659489307e2565ebfd17774642c9ae5d3c18068dc35170c7d58 -4cf4173f1baf98137fa249c81f3347e1dadd6b1ba0f50c3b64c1eab183a0 -937b0f7278eff101e5267fa6480da7d602844416490c2c2c7eb0d44ac8f4 -75cfd611db5ec268db07c0b3608825c3e12834a2b2efaf5e2723c5199c42 -6011cf22e64e4c0d31d563f321097935ea0c6fcbf5acd3748d90079f6ab8 -687288dc55df29fe7958f566b27b73e2ea30747247f7a2b2add0602c7d64 -d23f52e7c96748e6a54ee8c4629b2aab8882169653f0ba7f05236bf14364 -244720f3259cbed73a318b29e4a9305deb65a2c9dec8a9d0f9a9f6fae541 -83e0f4b9a9a567057a1794945168dc23cec25d1c02ea9242c9fb6d8fc11e -e8874bd80a5226373ae87cea91853d0625c777ceb1f5a6f3debcf2f75a61 -460c7b4067f568ecd01f62901ade8bf8fbc5db9c6720420496f0cb48a002 -99870773c2e7b12e83987a5d0290d9bbf589ac889bf7d4334a5147187a7f -71008f216ce917ca4cfba5347078f354897fd87ac48af6a6c62711d2eb3a -5882bf3b32c0f1bfda976f850c9dcb97170e78c229a27fd5e292d161ece9 -a8c47a223cbdc28e24f79f6429c72b5752a08f917feda941582c36d9acb5 -748c86072858d053170fdbf708971a0bd5a8d8034ec769cb72ea88eb5cd7 -49f35be6ee5e9b5df6021926cae9dac3f5ec2b33680b12e95fd4ecbf28eb -a0503c10c6f2be6c7c47e9d66a0fae6038441c50e6447892f4aaf0a25ccd -952c2e8b201bb479099f16fc4903993ac18d4667c84c124685ae7648a826 -6bc1701cc600964fdcc01258a72104a0e5e9996b34c2691a66fa20f48d7c -2522333dfdabf3785f37dd9b021e8ee29fa10f76f43d5f935996cbf9d98d -92d0a84ce65613f7c4a5052f4c408bf10679fc28a4a9ff848d9e0c4976bb -dfdfb78bb934cd72434db596cb49e199f386a0bda69449ce2e11e3a4f53d -be134c6d7fe452a0927cf6a9a15b2406f8bd354adcde0ce136378baa565f -b9c51a03b1fbe1e166a1f92af26bd9f072250aaa6596a236ba2d5a200c90 -a760ca050421abc78223b2e8b2eea958ab23084fa1947574e846e48aeb12 -26cebb8b5a92089e9ea771557599e2fff44d75bcf600e76ae7289ba98cf3 -98208c5104562834f568ebd62801b988b0a9fdf132b6564566103b3d2d8e -6a099b7fbad8a13b8cd7f6729bb6651fc1019e66c4bd6ff27410bd5cdae7 -4010bd68b066bffdb4fd5e3dd9cf7e1a1353f7a4c5157e3ad508f4ca0259 -9761b7cdd6a81b3560b8765be3b0432fe4c25dcb4001b00c7fa62874f681 -ed22127dc3974605a05be8d8fcf9701f859ffce4dc598091891ab7596ac3 -4cd851ecfd2dbbaa2f99dac376f7bb40703fd0700d7499a7c24726bdc9bb -3b88c6a82e52686c1ee945d8825092bc81848a08722ac5a1d24353f95ec8 -18f3fa487d9600318091b0ae9874b42bb3cb683a2518b18cc1bd86c6e5e8 -3d37c14ef4fe0c77b03a3314995b1e7c1066b98c4375bd1fc5fadee1b024 -7ece4f95a0f59978d543910deb2e5761632c74c508269c4e4b9e315bda02 -975dc771fc30c8164b9df9172a4e571d8ca578cd2aaeaa0dd083e74cdc2e -d938b984b96d76a64b8c5fd12e63220bbac41e5bcd5ccb6b84bdbf6a02d5 -934ac50c654c0853209a6758bcdf560e53566d78987484bb6672ebe93f22 -dcba14e3acc132a2d9ae837adde04d8b16 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -cleartomark -%%BeginResource: procset Altsys_header 4 0 -userdict begin /AltsysDict 245 dict def end -AltsysDict begin -/bdf{bind def}bind def -/xdf{exch def}bdf -/defed{where{pop true}{false}ifelse}bdf -/ndf{1 index where{pop pop pop}{dup xcheck{bind}if def}ifelse}bdf -/d{setdash}bdf -/h{closepath}bdf -/H{}bdf -/J{setlinecap}bdf -/j{setlinejoin}bdf -/M{setmiterlimit}bdf -/n{newpath}bdf -/N{newpath}bdf -/q{gsave}bdf -/Q{grestore}bdf -/w{setlinewidth}bdf -/sepdef{ - dup where not - { -AltsysSepDict - } - if - 3 1 roll exch put -}bdf -/st{settransfer}bdf -/colorimage defed /_rci xdf -/_NXLevel2 defed { - _NXLevel2 not { -/colorimage where { -userdict eq { -/_rci false def -} if -} if - } if -} if -/md defed{ - md type /dicttype eq { -/colorimage where { -md eq { -/_rci false def -}if -}if -/settransfer where { -md eq { -/st systemdict /settransfer get def -}if -}if - }if -}if -/setstrokeadjust defed -{ - true setstrokeadjust - /C{curveto}bdf - /L{lineto}bdf - /m{moveto}bdf -} -{ - /dr{transform .25 sub round .25 add -exch .25 sub round .25 add exch itransform}bdf - /C{dr curveto}bdf - /L{dr lineto}bdf - /m{dr moveto}bdf - /setstrokeadjust{pop}bdf -}ifelse -/rectstroke defed /xt xdf -xt {/yt save def} if -/privrectpath { - 4 -2 roll m - dtransform round exch round exch idtransform - 2 copy 0 lt exch 0 lt xor - {dup 0 exch rlineto exch 0 rlineto neg 0 exch rlineto} - {exch dup 0 rlineto exch 0 exch rlineto neg 0 rlineto} - ifelse - closepath -}bdf -/rectclip{newpath privrectpath clip newpath}def -/rectfill{gsave newpath privrectpath fill grestore}def -/rectstroke{gsave newpath privrectpath stroke grestore}def -xt {yt restore} if -/_fonthacksave false def -/currentpacking defed -{ - /_bfh {/_fonthacksave currentpacking def false setpacking} bdf - /_efh {_fonthacksave setpacking} bdf -} -{ - /_bfh {} bdf - /_efh {} bdf -}ifelse -/packedarray{array astore readonly}ndf -/` -{ - false setoverprint - - - /-save0- save def - 5 index concat - pop - storerect left bottom width height rectclip - pop - - /dict_count countdictstack def - /op_count count 1 sub def - userdict begin - - /showpage {} def - - 0 setgray 0 setlinecap 1 setlinewidth - 0 setlinejoin 10 setmiterlimit [] 0 setdash newpath - -} bdf -/currentpacking defed{true setpacking}if -/min{2 copy gt{exch}if pop}bdf -/max{2 copy lt{exch}if pop}bdf -/xformfont { currentfont exch makefont setfont } bdf -/fhnumcolors 1 - statusdict begin -/processcolors defed -{ -pop processcolors -} -{ -/deviceinfo defed { -deviceinfo /Colors known { -pop deviceinfo /Colors get -} if -} if -} ifelse - end -def -/printerRes - gsave - matrix defaultmatrix setmatrix - 72 72 dtransform - abs exch abs - max - grestore - def -/graycalcs -[ - {Angle Frequency} - {GrayAngle GrayFrequency} - {0 Width Height matrix defaultmatrix idtransform -dup mul exch dup mul add sqrt 72 exch div} - {0 GrayWidth GrayHeight matrix defaultmatrix idtransform -dup mul exch dup mul add sqrt 72 exch div} -] def -/calcgraysteps { - forcemaxsteps - { -maxsteps - } - { -/currenthalftone defed -{currenthalftone /dicttype eq}{false}ifelse -{ -currenthalftone begin -HalftoneType 4 le -{graycalcs HalftoneType 1 sub get exec} -{ -HalftoneType 5 eq -{ -Default begin -{graycalcs HalftoneType 1 sub get exec} -end -} -{0 60} -ifelse -} -ifelse -end -} -{ -currentscreen pop exch -} -ifelse - -printerRes 300 max exch div exch -2 copy -sin mul round dup mul -3 1 roll -cos mul round dup mul -add 1 add -dup maxsteps gt {pop maxsteps} if - } - ifelse -} bdf -/nextrelease defed { - /languagelevel defed not { -/framebuffer defed { -0 40 string framebuffer 9 1 roll 8 {pop} repeat -dup 516 eq exch 520 eq or -{ -/fhnumcolors 3 def -/currentscreen {60 0 {pop pop 1}}bdf -/calcgraysteps {maxsteps} bdf -}if -}if - }if -}if -fhnumcolors 1 ne { - /calcgraysteps {maxsteps} bdf -} if -/currentpagedevice defed { - - - currentpagedevice /PreRenderingEnhance known - { -currentpagedevice /PreRenderingEnhance get -{ -/calcgraysteps -{ -forcemaxsteps -{maxsteps} -{256 maxsteps min} -ifelse -} def -} if - } if -} if -/gradfrequency 144 def -printerRes 1000 lt { - /gradfrequency 72 def -} if -/adjnumsteps { - - dup dtransform abs exch abs max - - printerRes div - - gradfrequency mul - round - 5 max - min -}bdf -/goodsep { - spots exch get 4 get dup sepname eq exch (_vc_Registration) eq or -}bdf -/BeginGradation defed -{/bb{BeginGradation}bdf} -{/bb{}bdf} -ifelse -/EndGradation defed -{/eb{EndGradation}bdf} -{/eb{}bdf} -ifelse -/bottom -0 def -/delta -0 def -/frac -0 def -/height -0 def -/left -0 def -/numsteps1 -0 def -/radius -0 def -/right -0 def -/top -0 def -/width -0 def -/xt -0 def -/yt -0 def -/df currentflat def -/tempstr 1 string def -/clipflatness currentflat def -/inverted? - 0 currenttransfer exec .5 ge def -/tc1 [0 0 0 1] def -/tc2 [0 0 0 1] def -/storerect{/top xdf /right xdf /bottom xdf /left xdf -/width right left sub def /height top bottom sub def}bdf -/concatprocs{ - systemdict /packedarray known - {dup type /packedarraytype eq 2 index type /packedarraytype eq or}{false}ifelse - { -/proc2 exch cvlit def /proc1 exch cvlit def -proc1 aload pop proc2 aload pop -proc1 length proc2 length add packedarray cvx - } - { -/proc2 exch cvlit def /proc1 exch cvlit def -/newproc proc1 length proc2 length add array def -newproc 0 proc1 putinterval newproc proc1 length proc2 putinterval -newproc cvx - }ifelse -}bdf -/i{dup 0 eq - {pop df dup} - {dup} ifelse - /clipflatness xdf setflat -}bdf -version cvr 38.0 le -{/setrgbcolor{ -currenttransfer exec 3 1 roll -currenttransfer exec 3 1 roll -currenttransfer exec 3 1 roll -setrgbcolor}bdf}if -/vms {/vmsv save def} bdf -/vmr {vmsv restore} bdf -/vmrs{vmsv restore /vmsv save def}bdf -/eomode{ - {/filler /eofill load def /clipper /eoclip load def} - {/filler /fill load def /clipper /clip load def} - ifelse -}bdf -/normtaper{}bdf -/logtaper{9 mul 1 add log}bdf -/CD{ - /NF exch def - { -exch dup -/FID ne 1 index/UniqueID ne and -{exch NF 3 1 roll put} -{pop pop} -ifelse - }forall - NF -}bdf -/MN{ - 1 index length - /Len exch def - dup length Len add - string dup - Len - 4 -1 roll - putinterval - dup - 0 - 4 -1 roll - putinterval -}bdf -/RC{4 -1 roll /ourvec xdf 256 string cvs(|______)anchorsearch - {1 index MN cvn/NewN exch def cvn - findfont dup maxlength dict CD dup/FontName NewN put dup - /Encoding ourvec put NewN exch definefont pop}{pop}ifelse}bdf -/RF{ - dup - FontDirectory exch - known - {pop 3 -1 roll pop} - {RC} - ifelse -}bdf -/FF{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known - {exch pop findfont 3 -1 roll pop} - {pop dup findfont dup maxlength dict CD dup dup - /Encoding exch /Encoding get 256 array copy 7 -1 roll - {3 -1 roll dup 4 -2 roll put}forall put definefont} - ifelse}bdf -/RFJ{ - dup - FontDirectory exch - known - {pop 3 -1 roll pop - FontDirectory /Ryumin-Light-83pv-RKSJ-H known - {pop pop /Ryumin-Light-83pv-RKSJ-H dup}if - } - {RC} - ifelse -}bdf -/FFJ{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known - {exch pop findfont 3 -1 roll pop} - {pop -dup FontDirectory exch known not - {FontDirectory /Ryumin-Light-83pv-RKSJ-H known -{pop /Ryumin-Light-83pv-RKSJ-H}if - }if - dup findfont dup maxlength dict CD dup dup - /Encoding exch /Encoding get 256 array copy 7 -1 roll - {3 -1 roll dup 4 -2 roll put}forall put definefont} - ifelse}bdf -/fps{ - currentflat - exch - dup 0 le{pop 1}if - { -dup setflat 3 index stopped -{1.3 mul dup 3 index gt{pop setflat pop pop stop}if} -{exit} -ifelse - }loop - pop setflat pop pop -}bdf -/fp{100 currentflat fps}bdf -/clipper{clip}bdf -/W{/clipper load 100 clipflatness dup setflat fps}bdf -userdict begin /BDFontDict 29 dict def end -BDFontDict begin -/bu{}def -/bn{}def -/setTxMode{av 70 ge{pop}if pop}def -/gm{m}def -/show{pop}def -/gr{pop}def -/fnt{pop pop pop}def -/fs{pop}def -/fz{pop}def -/lin{pop pop}def -/:M {pop pop} def -/sf {pop} def -/S {pop} def -/@b {pop pop pop pop pop pop pop pop} def -/_bdsave /save load def -/_bdrestore /restore load def -/save { dup /fontsave eq {null} {_bdsave} ifelse } def -/restore { dup null eq { pop } { _bdrestore } ifelse } def -/fontsave null def -end -/MacVec 256 array def -MacVec 0 /Helvetica findfont -/Encoding get 0 128 getinterval putinterval -MacVec 127 /DEL put MacVec 16#27 /quotesingle put MacVec 16#60 /grave put -/NUL/SOH/STX/ETX/EOT/ENQ/ACK/BEL/BS/HT/LF/VT/FF/CR/SO/SI -/DLE/DC1/DC2/DC3/DC4/NAK/SYN/ETB/CAN/EM/SUB/ESC/FS/GS/RS/US -MacVec 0 32 getinterval astore pop -/Adieresis/Aring/Ccedilla/Eacute/Ntilde/Odieresis/Udieresis/aacute -/agrave/acircumflex/adieresis/atilde/aring/ccedilla/eacute/egrave -/ecircumflex/edieresis/iacute/igrave/icircumflex/idieresis/ntilde/oacute -/ograve/ocircumflex/odieresis/otilde/uacute/ugrave/ucircumflex/udieresis -/dagger/degree/cent/sterling/section/bullet/paragraph/germandbls -/registered/copyright/trademark/acute/dieresis/notequal/AE/Oslash -/infinity/plusminus/lessequal/greaterequal/yen/mu/partialdiff/summation -/product/pi/integral/ordfeminine/ordmasculine/Omega/ae/oslash -/questiondown/exclamdown/logicalnot/radical/florin/approxequal/Delta/guillemotleft -/guillemotright/ellipsis/nbspace/Agrave/Atilde/Otilde/OE/oe -/endash/emdash/quotedblleft/quotedblright/quoteleft/quoteright/divide/lozenge -/ydieresis/Ydieresis/fraction/currency/guilsinglleft/guilsinglright/fi/fl -/daggerdbl/periodcentered/quotesinglbase/quotedblbase -/perthousand/Acircumflex/Ecircumflex/Aacute -/Edieresis/Egrave/Iacute/Icircumflex/Idieresis/Igrave/Oacute/Ocircumflex -/apple/Ograve/Uacute/Ucircumflex/Ugrave/dotlessi/circumflex/tilde -/macron/breve/dotaccent/ring/cedilla/hungarumlaut/ogonek/caron -MacVec 128 128 getinterval astore pop -end %. AltsysDict -%%EndResource -%%EndProlog -%%BeginSetup -AltsysDict begin -_bfh -%%IncludeResource: font Symbol -_efh -0 dict dup begin -end -/f0 /Symbol FF def -_bfh -%%IncludeResource: font ZapfHumanist601BT-Bold -_efh -0 dict dup begin -end -/f1 /ZapfHumanist601BT-Bold FF def -end %. AltsysDict -%%EndSetup -AltsysDict begin -/onlyk4{false}ndf -/ccmyk{dup 5 -1 roll sub 0 max exch}ndf -/cmyk2gray{ - 4 -1 roll 0.3 mul 4 -1 roll 0.59 mul 4 -1 roll 0.11 mul - add add add 1 min neg 1 add -}bdf -/setcmykcolor{1 exch sub ccmyk ccmyk ccmyk pop setrgbcolor}ndf -/maxcolor { - max max max -} ndf -/maxspot { - pop -} ndf -/setcmykcoloroverprint{4{dup -1 eq{pop 0}if 4 1 roll}repeat setcmykcolor}ndf -/findcmykcustomcolor{5 packedarray}ndf -/setcustomcolor{exch aload pop pop 4{4 index mul 4 1 roll}repeat setcmykcolor pop}ndf -/setseparationgray{setgray}ndf -/setoverprint{pop}ndf -/currentoverprint false ndf -/cmykbufs2gray{ - 0 1 2 index length 1 sub - { -4 index 1 index get 0.3 mul -4 index 2 index get 0.59 mul -4 index 3 index get 0.11 mul -4 index 4 index get -add add add cvi 255 min -255 exch sub -2 index 3 1 roll put - }for - 4 1 roll pop pop pop -}bdf -/colorimage{ - pop pop - [ -5 -1 roll/exec cvx -6 -1 roll/exec cvx -7 -1 roll/exec cvx -8 -1 roll/exec cvx -/cmykbufs2gray cvx - ]cvx - image -} -%. version 47.1 on Linotronic of Postscript defines colorimage incorrectly (rgb model only) -version cvr 47.1 le -statusdict /product get (Lino) anchorsearch{pop pop true}{pop false}ifelse -and{userdict begin bdf end}{ndf}ifelse -fhnumcolors 1 ne {/yt save def} if -/customcolorimage{ - aload pop - (_vc_Registration) eq - { -pop pop pop pop separationimage - } - { -/ik xdf /iy xdf /im xdf /ic xdf -ic im iy ik cmyk2gray /xt xdf -currenttransfer -{dup 1.0 exch sub xt mul add}concatprocs -st -image - } - ifelse -}ndf -fhnumcolors 1 ne {yt restore} if -fhnumcolors 3 ne {/yt save def} if -/customcolorimage{ - aload pop - (_vc_Registration) eq - { -pop pop pop pop separationimage - } - { -/ik xdf /iy xdf /im xdf /ic xdf -1.0 dup ic ik add min sub -1.0 dup im ik add min sub -1.0 dup iy ik add min sub -/ic xdf /iy xdf /im xdf -currentcolortransfer -4 1 roll -{dup 1.0 exch sub ic mul add}concatprocs 4 1 roll -{dup 1.0 exch sub iy mul add}concatprocs 4 1 roll -{dup 1.0 exch sub im mul add}concatprocs 4 1 roll -setcolortransfer -{/dummy xdf dummy}concatprocs{dummy}{dummy}true 3 colorimage - } - ifelse -}ndf -fhnumcolors 3 ne {yt restore} if -fhnumcolors 4 ne {/yt save def} if -/customcolorimage{ - aload pop - (_vc_Registration) eq - { -pop pop pop pop separationimage - } - { -/ik xdf /iy xdf /im xdf /ic xdf -currentcolortransfer -{1.0 exch sub ik mul ik sub 1 add}concatprocs 4 1 roll -{1.0 exch sub iy mul iy sub 1 add}concatprocs 4 1 roll -{1.0 exch sub im mul im sub 1 add}concatprocs 4 1 roll -{1.0 exch sub ic mul ic sub 1 add}concatprocs 4 1 roll -setcolortransfer -{/dummy xdf dummy}concatprocs{dummy}{dummy}{dummy} -true 4 colorimage - } - ifelse -}ndf -fhnumcolors 4 ne {yt restore} if -/separationimage{image}ndf -/newcmykcustomcolor{6 packedarray}ndf -/inkoverprint false ndf -/setinkoverprint{pop}ndf -/setspotcolor { - spots exch get - dup 4 get (_vc_Registration) eq - {pop 1 exch sub setseparationgray} - {0 5 getinterval exch setcustomcolor} - ifelse -}ndf -/currentcolortransfer{currenttransfer dup dup dup}ndf -/setcolortransfer{st pop pop pop}ndf -/fas{}ndf -/sas{}ndf -/fhsetspreadsize{pop}ndf -/filler{fill}bdf -/F{gsave {filler}fp grestore}bdf -/f{closepath F}bdf -/S{gsave {stroke}fp grestore}bdf -/s{closepath S}bdf -/bc4 [0 0 0 0] def -/_lfp4 { - /iosv inkoverprint def - /cosv currentoverprint def - /yt xdf - /xt xdf - /ang xdf - storerect - /taperfcn xdf - /k2 xdf /y2 xdf /m2 xdf /c2 xdf - /k1 xdf /y1 xdf /m1 xdf /c1 xdf - c1 c2 sub abs - m1 m2 sub abs - y1 y2 sub abs - k1 k2 sub abs - maxcolor - calcgraysteps mul abs round - height abs adjnumsteps - dup 2 lt {pop 1} if - 1 sub /numsteps1 xdf - currentflat mark - currentflat clipflatness - /delta top bottom sub numsteps1 1 add div def - /right right left sub def - /botsv top delta sub def - { -{ -W -xt yt translate -ang rotate -xt neg yt neg translate -dup setflat -/bottom botsv def -0 1 numsteps1 -{ -numsteps1 dup 0 eq {pop 0.5 } { div } ifelse -taperfcn /frac xdf -bc4 0 c2 c1 sub frac mul c1 add put -bc4 1 m2 m1 sub frac mul m1 add put -bc4 2 y2 y1 sub frac mul y1 add put -bc4 3 k2 k1 sub frac mul k1 add put -bc4 vc -1 index setflat -{ -mark {newpath left bottom right delta rectfill}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -/bottom bottom delta sub def -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/bcs [0 0] def -/_lfs4 { - /iosv inkoverprint def - /cosv currentoverprint def - /yt xdf - /xt xdf - /ang xdf - storerect - /taperfcn xdf - /tint2 xdf - /tint1 xdf - bcs exch 1 exch put - tint1 tint2 sub abs - bcs 1 get maxspot - calcgraysteps mul abs round - height abs adjnumsteps - dup 2 lt {pop 2} if - 1 sub /numsteps1 xdf - currentflat mark - currentflat clipflatness - /delta top bottom sub numsteps1 1 add div def - /right right left sub def - /botsv top delta sub def - { -{ -W -xt yt translate -ang rotate -xt neg yt neg translate -dup setflat -/bottom botsv def -0 1 numsteps1 -{ -numsteps1 div taperfcn /frac xdf -bcs 0 -1.0 tint2 tint1 sub frac mul tint1 add sub -put bcs vc -1 index setflat -{ -mark {newpath left bottom right delta rectfill}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -/bottom bottom delta sub def -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/_rfs4 { - /iosv inkoverprint def - /cosv currentoverprint def - /tint2 xdf - /tint1 xdf - bcs exch 1 exch put - /radius xdf - /yt xdf - /xt xdf - tint1 tint2 sub abs - bcs 1 get maxspot - calcgraysteps mul abs round - radius abs adjnumsteps - dup 2 lt {pop 2} if - 1 sub /numsteps1 xdf - radius numsteps1 div 2 div /halfstep xdf - currentflat mark - currentflat clipflatness - { -{ -dup setflat -W -0 1 numsteps1 -{ -dup /radindex xdf -numsteps1 div /frac xdf -bcs 0 -tint2 tint1 sub frac mul tint1 add -put bcs vc -1 index setflat -{ -newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 -{ arc -radindex numsteps1 ne -{ -xt yt -radindex 1 add numsteps1 -div 1 exch sub -radius mul halfstep add -dup xt add yt moveto -360 0 arcn -} if -fill -}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/_rfp4 { - /iosv inkoverprint def - /cosv currentoverprint def - /k2 xdf /y2 xdf /m2 xdf /c2 xdf - /k1 xdf /y1 xdf /m1 xdf /c1 xdf - /radius xdf - /yt xdf - /xt xdf - c1 c2 sub abs - m1 m2 sub abs - y1 y2 sub abs - k1 k2 sub abs - maxcolor - calcgraysteps mul abs round - radius abs adjnumsteps - dup 2 lt {pop 1} if - 1 sub /numsteps1 xdf - radius numsteps1 dup 0 eq {pop} {div} ifelse - 2 div /halfstep xdf - currentflat mark - currentflat clipflatness - { -{ -dup setflat -W -0 1 numsteps1 -{ -dup /radindex xdf -numsteps1 dup 0 eq {pop 0.5 } { div } ifelse -/frac xdf -bc4 0 c2 c1 sub frac mul c1 add put -bc4 1 m2 m1 sub frac mul m1 add put -bc4 2 y2 y1 sub frac mul y1 add put -bc4 3 k2 k1 sub frac mul k1 add put -bc4 vc -1 index setflat -{ -newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 -{ arc -radindex numsteps1 ne -{ -xt yt -radindex 1 add -numsteps1 dup 0 eq {pop} {div} ifelse -1 exch sub -radius mul halfstep add -dup xt add yt moveto -360 0 arcn -} if -fill -}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/lfp4{_lfp4}ndf -/lfs4{_lfs4}ndf -/rfs4{_rfs4}ndf -/rfp4{_rfp4}ndf -/cvc [0 0 0 1] def -/vc{ - AltsysDict /cvc 2 index put - aload length 4 eq - {setcmykcolor} - {setspotcolor} - ifelse -}bdf -/origmtx matrix currentmatrix def -/ImMatrix matrix currentmatrix def -0 setseparationgray -/imgr {1692 1570.1102 2287.2756 2412 } def -/bleed 0 def -/clpr {1692 1570.1102 2287.2756 2412 } def -/xs 1 def -/ys 1 def -/botx 0 def -/overlap 0 def -/wdist 18 def -0 2 mul fhsetspreadsize -0 0 ne {/df 0 def /clipflatness 0 def} if -/maxsteps 256 def -/forcemaxsteps false def -vms --1845 -1956 translate -/currentpacking defed{false setpacking}if -/spots[ -1 0 0 0 (Process Cyan) false newcmykcustomcolor -0 1 0 0 (Process Magenta) false newcmykcustomcolor -0 0 1 0 (Process Yellow) false newcmykcustomcolor -0 0 0 1 (Process Black) false newcmykcustomcolor -]def -/textopf false def -/curtextmtx{}def -/otw .25 def -/msf{dup/curtextmtx xdf makefont setfont}bdf -/makesetfont/msf load def -/curtextheight{.707104 .707104 curtextmtx dtransform - dup mul exch dup mul add sqrt}bdf -/ta2{ -tempstr 2 index gsave exec grestore -cwidth cheight rmoveto -4 index eq{5 index 5 index rmoveto}if -2 index 2 index rmoveto -}bdf -/ta{exch systemdict/cshow known -{{/cheight xdf/cwidth xdf tempstr 0 2 index put ta2}exch cshow} -{{tempstr 0 2 index put tempstr stringwidth/cheight xdf/cwidth xdf ta2}forall} -ifelse 6{pop}repeat}bdf -/sts{/textopf currentoverprint def vc setoverprint -/ts{awidthshow}def exec textopf setoverprint}bdf -/stol{/xt currentlinewidth def - setlinewidth vc newpath - /ts{{false charpath stroke}ta}def exec - xt setlinewidth}bdf - -/strk{/textopf currentoverprint def vc setoverprint - /ts{{false charpath stroke}ta}def exec - textopf setoverprint - }bdf -n -[] 0 d -3.863708 M -1 w -0 j -0 J -false setoverprint -0 i -false eomode -[0 0 0 1] vc -vms -%white border -- disabled -%1845.2293 2127.8588 m -%2045.9437 2127.8588 L -%2045.9437 1956.1412 L -%1845.2293 1956.1412 L -%1845.2293 2127.8588 L -%0.1417 w -%2 J -%2 M -%[0 0 0 0] vc -%s -n -1950.8 2097.2 m -1958.8 2092.5 1967.3 2089 1975.5 2084.9 C -1976.7 2083.5 1976.1 2081.5 1976.7 2079.9 C -1979.6 2081.1 1981.6 2086.8 1985.3 2084 C -1993.4 2079.3 2001.8 2075.8 2010 2071.7 C -2010.5 2071.5 2010.5 2071.1 2010.8 2070.8 C -2011.2 2064.3 2010.9 2057.5 2011 2050.8 C -2015.8 2046.9 2022.2 2046.2 2026.6 2041.7 C -2026.5 2032.5 2026.8 2022.9 2026.4 2014.1 C -2020.4 2008.3 2015 2002.4 2008.8 1997.1 C -2003.8 1996.8 2000.7 2001.2 1996.1 2002.1 C -1995.2 1996.4 1996.9 1990.5 1995.6 1984.8 C -1989.9 1979 1984.5 1973.9 1978.8 1967.8 C -1977.7 1968.6 1976 1967.6 1974.5 1968.3 C -1967.4 1972.5 1960.1 1976.1 1952.7 1979.3 C -1946.8 1976.3 1943.4 1970.7 1938.5 1966.1 C -1933.9 1966.5 1929.4 1968.8 1925.1 1970.7 C -1917.2 1978.2 1906 1977.9 1897.2 1983.4 C -1893.2 1985.6 1889.4 1988.6 1885 1990.1 C -1884.6 1990.6 1883.9 1991 1883.8 1991.6 C -1883.7 2000.4 1884 2009.9 1883.6 2018.9 C -1887.7 2024 1893.2 2028.8 1898 2033.8 C -1899.1 2035.5 1900.9 2036.8 1902.5 2037.9 C -1903.9 2037.3 1905.2 2036.6 1906.4 2035.5 C -1906.3 2039.7 1906.5 2044.6 1906.1 2048.9 C -1906.3 2049.6 1906.7 2050.2 1907.1 2050.8 C -1913.4 2056 1918.5 2062.7 1924.8 2068.1 C -1926.6 2067.9 1928 2066.9 1929.4 2066 C -1930.2 2071 1927.7 2077.1 1930.6 2081.6 C -1936.6 2086.9 1941.5 2092.9 1947.9 2097.9 C -1949 2098.1 1949.9 2097.5 1950.8 2097.2 C -[0 0 0 0.18] vc -f -0.4 w -S -n -1975.2 2084.7 m -1976.6 2083.4 1975.7 2081.1 1976 2079.4 C -1979.3 2079.5 1980.9 2086.2 1984.8 2084 C -1992.9 2078.9 2001.7 2075.6 2010 2071.2 C -2011 2064.6 2010.2 2057.3 2010.8 2050.6 C -2015.4 2046.9 2021.1 2045.9 2025.9 2042.4 C -2026.5 2033.2 2026.8 2022.9 2025.6 2013.9 C -2020.5 2008.1 2014.5 2003.1 2009.3 1997.6 C -2004.1 1996.7 2000.7 2001.6 1995.9 2002.6 C -1995.2 1996.7 1996.3 1990.2 1994.9 1984.6 C -1989.8 1978.7 1983.6 1973.7 1978.4 1968 C -1977.3 1969.3 1976 1967.6 1974.8 1968.5 C -1967.7 1972.7 1960.4 1976.3 1952.9 1979.6 C -1946.5 1976.9 1943.1 1970.5 1937.8 1966.1 C -1928.3 1968.2 1920.6 1974.8 1911.6 1978.4 C -1901.9 1979.7 1893.9 1986.6 1885 1990.6 C -1884.3 1991 1884.3 1991.7 1884 1992.3 C -1884.5 2001 1884.2 2011 1884.3 2019.9 C -1890.9 2025.3 1895.9 2031.9 1902.3 2037.4 C -1904.2 2037.9 1905.6 2034.2 1906.8 2035.7 C -1907.4 2040.9 1905.7 2046.1 1907.3 2050.8 C -1913.6 2056.2 1919.2 2062.6 1925.1 2067.9 C -1926.9 2067.8 1928 2066.3 1929.6 2065.7 C -1929.9 2070.5 1929.2 2076 1930.1 2080.8 C -1936.5 2086.1 1941.6 2092.8 1948.4 2097.6 C -1957.3 2093.3 1966.2 2088.8 1975.2 2084.7 C -[0 0 0 0] vc -f -S -n -1954.8 2093.8 m -1961.6 2090.5 1968.2 2087 1975 2084 C -1975 2082.8 1975.6 2080.9 1974.8 2080.6 C -1974.3 2075.2 1974.6 2069.6 1974.5 2064 C -1977.5 2059.7 1984.5 2060 1988.9 2056.4 C -1989.5 2055.5 1990.5 2055.3 1990.8 2054.4 C -1991.1 2045.7 1991.4 2036.1 1990.6 2027.8 C -1990.7 2026.6 1992 2027.3 1992.8 2027.1 C -1997 2032.4 2002.6 2037.8 2007.6 2042.2 C -2008.7 2042.3 2007.8 2040.6 2007.4 2040 C -2002.3 2035.6 1997.5 2030 1992.8 2025.2 C -1991.6 2024.7 1990.8 2024.9 1990.1 2025.4 C -1989.4 2024.9 1988.1 2025.2 1987.2 2024.4 C -1987.1 2025.8 1988.3 2026.5 1989.4 2026.8 C -1989.4 2026.6 1989.3 2026.2 1989.6 2026.1 C -1989.9 2026.2 1989.9 2026.6 1989.9 2026.8 C -1989.8 2026.6 1990 2026.5 1990.1 2026.4 C -1990.2 2027 1991.1 2028.3 1990.1 2028 C -1989.9 2037.9 1990.5 2044.1 1989.6 2054.2 C -1985.9 2058 1979.7 2057.4 1976 2061.2 C -1974.5 2061.6 1975.2 2059.9 1974.5 2059.5 C -1973.9 2058 1975.6 2057.8 1975 2056.6 C -1974.5 2057.1 1974.6 2055.3 1973.6 2055.9 C -1971.9 2059.3 1974.7 2062.1 1973.1 2065.5 C -1973.1 2071.2 1972.9 2077 1973.3 2082.5 C -1967.7 2085.6 1962 2088 1956.3 2090.7 C -1953.9 2092.4 1951 2093 1948.6 2094.8 C -1943.7 2089.9 1937.9 2084.3 1933 2079.6 C -1931.3 2076.1 1933.2 2071.3 1932.3 2067.2 C -1931.3 2062.9 1933.3 2060.6 1932 2057.6 C -1932.7 2056.5 1930.9 2053.3 1933.2 2051.8 C -1936.8 2050.1 1940.1 2046.9 1944 2046.8 C -1946.3 2049.7 1949.3 2051.9 1952 2054.4 C -1954.5 2054.2 1956.4 2052.3 1958.7 2051.3 C -1960.8 2050 1963.2 2049 1965.6 2048.4 C -1968.3 2050.8 1970.7 2054.3 1973.6 2055.4 C -1973 2052.2 1969.7 2050.4 1967.6 2048.2 C -1967.1 2046.7 1968.8 2046.6 1969.5 2045.8 C -1972.8 2043.3 1980.6 2043.4 1979.3 2038.4 C -1979.4 2038.6 1979.2 2038.7 1979.1 2038.8 C -1978.7 2038.6 1978.9 2038.1 1978.8 2037.6 C -1978.9 2037.9 1978.7 2038 1978.6 2038.1 C -1978.2 2032.7 1978.4 2027.1 1978.4 2021.6 C -1979.3 2021.1 1980 2020.2 1981.5 2020.1 C -1983.5 2020.5 1984 2021.8 1985.1 2023.5 C -1985.7 2024 1987.4 2023.7 1986 2022.8 C -1984.7 2021.7 1983.3 2020.8 1983.9 2018.7 C -1987.2 2015.9 1993 2015.4 1994.9 2011.5 C -1992.2 2004.9 1999.3 2005.2 2002.1 2002.4 C -2005.9 2002.7 2004.8 1997.4 2009.1 1999 C -2011 1999.3 2010 2002.9 2012.7 2002.4 C -2010.2 2000.7 2009.4 1996.1 2005.5 1998.5 C -2002.1 2000.3 1999 2002.5 1995.4 2003.8 C -1995.2 2003.6 1994.9 2003.3 1994.7 2003.1 C -1994.3 1997 1995.6 1991.1 1994.4 1985.3 C -1994.3 1986 1993.8 1985 1994 1985.6 C -1993.8 1995.4 1994.4 2001.6 1993.5 2011.7 C -1989.7 2015.5 1983.6 2014.9 1979.8 2018.7 C -1978.3 2019.1 1979.1 2017.4 1978.4 2017 C -1977.8 2015.5 1979.4 2015.3 1978.8 2014.1 C -1978.4 2014.6 1978.5 2012.8 1977.4 2013.4 C -1975.8 2016.8 1978.5 2019.6 1976.9 2023 C -1977 2028.7 1976.7 2034.5 1977.2 2040 C -1971.6 2043.1 1965.8 2045.6 1960.1 2048.2 C -1957.7 2049.9 1954.8 2050.5 1952.4 2052.3 C -1947.6 2047.4 1941.8 2041.8 1936.8 2037.2 C -1935.2 2033.6 1937.1 2028.8 1936.1 2024.7 C -1935.1 2020.4 1937.1 2018.1 1935.9 2015.1 C -1936.5 2014.1 1934.7 2010.8 1937.1 2009.3 C -1944.4 2004.8 1952 2000.9 1959.9 1997.8 C -1963.9 1997 1963.9 2001.9 1966.8 2003.3 C -1970.3 2006.9 1973.7 2009.9 1976.9 2012.9 C -1977.9 2013 1977.1 2011.4 1976.7 2010.8 C -1971.6 2006.3 1966.8 2000.7 1962 1995.9 C -1960 1995.2 1960.1 1996.6 1958.2 1995.6 C -1957 1997 1955.1 1998.8 1953.2 1998 C -1951.7 1994.5 1954.1 1993.4 1952.9 1991.1 C -1952.1 1990.5 1953.3 1990.2 1953.2 1989.6 C -1954.2 1986.8 1950.9 1981.4 1954.4 1981.2 C -1954.7 1981.6 1954.7 1981.7 1955.1 1982 C -1961.9 1979.1 1967.6 1975 1974.3 1971.6 C -1974.7 1969.8 1976.7 1969.5 1978.4 1969.7 C -1980.3 1970 1979.3 1973.6 1982 1973.1 C -1975.8 1962.2 1968 1975.8 1960.8 1976.7 C -1956.9 1977.4 1953.3 1982.4 1949.1 1978.8 C -1946 1975.8 1941.2 1971 1939.5 1969.2 C -1938.5 1968.6 1938.9 1967.4 1937.8 1966.8 C -1928.7 1969.4 1920.6 1974.5 1912.4 1979.1 C -1904 1980 1896.6 1985 1889.3 1989.4 C -1887.9 1990.4 1885.1 1990.3 1885 1992.5 C -1885.4 2000.6 1885.2 2012.9 1885.2 2019.9 C -1886.1 2022 1889.7 2019.5 1888.4 2022.8 C -1889 2023.3 1889.8 2024.4 1890.3 2024 C -1891.2 2023.5 1891.8 2028.2 1893.4 2026.6 C -1894.2 2026.3 1893.9 2027.3 1894.4 2027.6 C -1893.4 2027.6 1894.7 2028.3 1894.1 2028.5 C -1894.4 2029.6 1896 2030 1896 2029.2 C -1896.2 2029 1896.3 2029 1896.5 2029.2 C -1896.8 2029.8 1897.3 2030 1897 2030.7 C -1896.5 2030.7 1896.9 2031.5 1897.2 2031.6 C -1898.3 2034 1899.5 2030.6 1899.6 2033.3 C -1898.5 2033 1899.6 2034.4 1900.1 2034.8 C -1901.3 2035.8 1903.2 2034.6 1902.5 2036.7 C -1904.4 2036.9 1906.1 2032.2 1907.6 2035.5 C -1907.5 2040.1 1907.7 2044.9 1907.3 2049.4 C -1908 2050.2 1908.3 2051.4 1909.5 2051.6 C -1910.1 2051.1 1911.6 2051.1 1911.4 2052.3 C -1909.7 2052.8 1912.4 2054 1912.6 2054.7 C -1913.4 2055.2 1913 2053.7 1913.6 2054.4 C -1913.6 2054.5 1913.6 2055.3 1913.6 2054.7 C -1913.7 2054.4 1913.9 2054.4 1914 2054.7 C -1914 2054.9 1914.1 2055.3 1913.8 2055.4 C -1913.7 2056 1915.2 2057.6 1916 2057.6 C -1915.9 2057.3 1916.1 2057.2 1916.2 2057.1 C -1917 2056.8 1916.7 2057.7 1917.2 2058 C -1917 2058.3 1916.7 2058.3 1916.4 2058.3 C -1917.1 2059 1917.3 2060.1 1918.4 2060.4 C -1918.1 2059.2 1919.1 2060.6 1919.1 2059.5 C -1919 2060.6 1920.6 2060.1 1919.8 2061.2 C -1919.6 2061.2 1919.3 2061.2 1919.1 2061.2 C -1919.6 2061.9 1921.4 2064.2 1921.5 2062.6 C -1922.4 2062.1 1921.6 2063.9 1922.2 2064.3 C -1922.9 2067.3 1926.1 2064.3 1925.6 2067.2 C -1927.2 2066.8 1928.4 2064.6 1930.1 2065.2 C -1931.8 2067.8 1931 2071.8 1930.8 2074.8 C -1930.6 2076.4 1930.1 2078.6 1930.6 2080.4 C -1936.6 2085.4 1941.8 2091.6 1948.1 2096.9 C -1950.7 2096.7 1952.6 2094.8 1954.8 2093.8 C -[0 0.33 0.33 0.99] vc -f -S -n -1989.4 2080.6 m -1996.1 2077.3 2002.7 2073.8 2009.6 2070.8 C -2009.6 2069.6 2010.2 2067.7 2009.3 2067.4 C -2008.9 2062 2009.1 2056.4 2009.1 2050.8 C -2012.3 2046.6 2019 2046.6 2023.5 2043.2 C -2024 2042.3 2025.1 2042.1 2025.4 2041.2 C -2025.3 2032.7 2025.6 2023.1 2025.2 2014.6 C -2025 2015.3 2024.6 2014.2 2024.7 2014.8 C -2024.5 2024.7 2025.1 2030.9 2024.2 2041 C -2020.4 2044.8 2014.3 2044.2 2010.5 2048 C -2009 2048.4 2009.8 2046.7 2009.1 2046.3 C -2008.5 2044.8 2010.2 2044.6 2009.6 2043.4 C -2009.1 2043.9 2009.2 2042.1 2008.1 2042.7 C -2006.5 2046.1 2009.3 2048.9 2007.6 2052.3 C -2007.7 2058 2007.5 2063.8 2007.9 2069.3 C -2002.3 2072.4 1996.5 2074.8 1990.8 2077.5 C -1988.4 2079.2 1985.6 2079.8 1983.2 2081.6 C -1980.5 2079 1977.9 2076.5 1975.5 2074.1 C -1975.5 2075.1 1975.5 2076.2 1975.5 2077.2 C -1977.8 2079.3 1980.3 2081.6 1982.7 2083.7 C -1985.3 2083.5 1987.1 2081.6 1989.4 2080.6 C -f -S -n -1930.1 2079.9 m -1931.1 2075.6 1929.2 2071.1 1930.8 2067.2 C -1930.3 2066.3 1930.1 2064.6 1928.7 2065.5 C -1927.7 2066.4 1926.5 2067 1925.3 2067.4 C -1924.5 2066.9 1925.6 2065.7 1924.4 2066 C -1924.2 2067.2 1923.6 2065.5 1923.2 2065.7 C -1922.3 2063.6 1917.8 2062.1 1919.6 2060.4 C -1919.3 2060.5 1919.2 2060.3 1919.1 2060.2 C -1919.7 2060.9 1918.2 2061 1917.6 2060.2 C -1917 2059.6 1916.1 2058.8 1916.4 2058 C -1915.5 2058 1917.4 2057.1 1915.7 2057.8 C -1914.8 2057.1 1913.4 2056.2 1913.3 2054.9 C -1913.1 2055.4 1911.3 2054.3 1910.9 2053.2 C -1910.7 2052.9 1910.2 2052.5 1910.7 2052.3 C -1911.1 2052.5 1910.9 2052 1910.9 2051.8 C -1910.5 2051.2 1909.9 2052.6 1909.2 2051.8 C -1908.2 2051.4 1907.8 2050.2 1907.1 2049.4 C -1907.5 2044.8 1907.3 2040 1907.3 2035.2 C -1905.3 2033 1902.8 2039.3 1902.3 2035.7 C -1899.6 2036 1898.4 2032.5 1896.3 2030.7 C -1895.7 2030.1 1897.5 2030 1896.3 2029.7 C -1896.3 2030.6 1895 2029.7 1894.4 2029.2 C -1892.9 2028.1 1894.2 2027.4 1893.6 2027.1 C -1892.1 2027.9 1891.7 2025.6 1890.8 2024.9 C -1891.1 2024.6 1889.1 2024.3 1888.4 2023 C -1887.5 2022.6 1888.2 2021.9 1888.1 2021.3 C -1886.7 2022 1885.2 2020.4 1884.8 2019.2 C -1884.8 2010 1884.6 2000.2 1885 1991.8 C -1886.9 1989.6 1889.9 1989.3 1892.2 1987.5 C -1898.3 1982.7 1905.6 1980.1 1912.8 1978.6 C -1921 1974.2 1928.8 1968.9 1937.8 1966.6 C -1939.8 1968.3 1938.8 1968.3 1940.4 1970 C -1945.4 1972.5 1947.6 1981.5 1954.6 1979.3 C -1952.3 1981 1950.4 1978.4 1948.6 1977.9 C -1945.1 1973.9 1941.1 1970.6 1938 1966.6 C -1928.4 1968.5 1920.6 1974.8 1911.9 1978.8 C -1907.1 1979.2 1902.6 1981.7 1898.2 1983.6 C -1893.9 1986 1889.9 1989 1885.5 1990.8 C -1884.9 1991.2 1884.8 1991.8 1884.5 1992.3 C -1884.9 2001.3 1884.7 2011.1 1884.8 2019.6 C -1890.6 2025 1896.5 2031.2 1902.3 2036.9 C -1904.6 2037.6 1905 2033 1907.3 2035.5 C -1907.2 2040.2 1907 2044.8 1907.1 2049.6 C -1913.6 2055.3 1918.4 2061.5 1925.1 2067.4 C -1927.3 2068.2 1929.6 2062.5 1930.6 2066.9 C -1929.7 2070.7 1930.3 2076 1930.1 2080.1 C -1935.6 2085.7 1941.9 2090.7 1947.2 2096.7 C -1942.2 2091.1 1935.5 2085.2 1930.1 2079.9 C -[0.18 0.18 0 0.78] vc -f -S -n -1930.8 2061.9 m -1930.3 2057.8 1931.8 2053.4 1931.1 2050.4 C -1931.3 2050.3 1931.7 2050.5 1931.6 2050.1 C -1933 2051.1 1934.4 2049.5 1935.9 2048.7 C -1937 2046.5 1939.5 2047.1 1941.2 2045.1 C -1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C -1934 2039.7 1934.5 2038.1 1933.7 2037.6 C -1934 2033.3 1933.1 2027.9 1934.4 2024.4 C -1934.3 2023.8 1933.9 2022.8 1933 2022.8 C -1931.6 2023.1 1930.5 2024.4 1929.2 2024.9 C -1928.4 2024.5 1929.8 2023.5 1928.7 2023.5 C -1927.7 2024.1 1926.2 2022.6 1925.6 2021.6 C -1926.9 2021.6 1924.8 2020.6 1925.6 2020.4 C -1924.7 2021.7 1923.9 2019.6 1923.2 2019.2 C -1923.3 2018.3 1923.8 2018.1 1923.2 2018 C -1922.9 2017.8 1922.9 2017.5 1922.9 2017.2 C -1922.8 2018.3 1921.3 2017.3 1920.3 2018 C -1916.6 2019.7 1913 2022.1 1910 2024.7 C -1910 2032.9 1910 2041.2 1910 2049.4 C -1915.4 2055.2 1920 2058.7 1925.3 2064.8 C -1927.2 2064 1929 2061.4 1930.8 2061.9 C -[0 0 0 0] vc -f -S -n -1907.6 2030.4 m -1907.5 2027.1 1906.4 2021.7 1908.5 2019.9 C -1908.8 2020.1 1908.9 2019 1909.2 2019.6 C -1910 2019.6 1912 2019.2 1913.1 2018.2 C -1913.7 2016.5 1920.2 2015.7 1917.4 2012.7 C -1918.2 2011.2 1917 2013.8 1917.2 2012 C -1916.9 2012.3 1916 2012.4 1915.2 2012 C -1912.5 2010.5 1916.6 2008.8 1913.6 2009.6 C -1912.6 2009.2 1911.1 2009 1910.9 2007.6 C -1911 1999.2 1911.8 1989.8 1911.2 1982.2 C -1910.1 1981.1 1908.8 1982.2 1907.6 1982.2 C -1900.8 1986.5 1893.2 1988.8 1887.2 1994.2 C -1887.2 2002.4 1887.2 2010.7 1887.2 2018.9 C -1892.6 2024.7 1897.2 2028.2 1902.5 2034.3 C -1904.3 2033.3 1906.2 2032.1 1907.6 2030.4 C -f -S -n -1910.7 2025.4 m -1912.7 2022.4 1916.7 2020.8 1919.8 2018.9 C -1920.2 2018.7 1920.6 2018.6 1921 2018.4 C -1925 2020 1927.4 2028.5 1932 2024.2 C -1932.3 2025 1932.5 2023.7 1932.8 2024.4 C -1932.8 2028 1932.8 2031.5 1932.8 2035 C -1931.9 2033.9 1932.5 2036.3 1932.3 2036.9 C -1933.2 2036.4 1932.5 2038.5 1933 2038.4 C -1933.1 2040.5 1935.6 2042.2 1936.6 2043.2 C -1936.2 2042.4 1935.1 2040.8 1933.7 2040.3 C -1932.2 2034.4 1933.8 2029.8 1933 2023.2 C -1931.1 2024.9 1928.4 2026.4 1926.5 2023.5 C -1925.1 2021.6 1923 2019.8 1921.5 2018.2 C -1917.8 2018.9 1915.2 2022.5 1911.6 2023.5 C -1910.8 2023.8 1911.2 2024.7 1910.4 2025.2 C -1910.9 2031.8 1910.6 2039.1 1910.7 2045.6 C -1910.1 2048 1910.7 2045.9 1911.2 2044.8 C -1910.6 2038.5 1911.2 2031.8 1910.7 2025.4 C -[0.07 0.06 0 0.58] vc -f -S -n -1910.7 2048.9 m -1910.3 2047.4 1911.3 2046.5 1911.6 2045.3 C -1912.9 2045.3 1913.9 2047.1 1915.2 2045.8 C -1915.2 2044.9 1916.6 2043.3 1917.2 2042.9 C -1918.7 2042.9 1919.4 2044.4 1920.5 2043.2 C -1921.2 2042.2 1921.4 2040.9 1922.4 2040.3 C -1924.5 2040.3 1925.7 2040.9 1926.8 2039.6 C -1927.1 2037.9 1926.8 2038.1 1927.7 2037.6 C -1929 2037.5 1930.4 2037 1931.6 2037.2 C -1932.3 2038.2 1933.1 2038.7 1932.8 2040.3 C -1935 2041.8 1935.9 2043.8 1938.5 2044.8 C -1938.6 2045 1938.3 2045.5 1938.8 2045.3 C -1939.1 2042.9 1935.4 2044.2 1935.4 2042.2 C -1932.1 2040.8 1932.8 2037.2 1932 2034.8 C -1932.3 2034 1932.7 2035.4 1932.5 2034.8 C -1931.3 2031.8 1935.5 2020.1 1928.9 2025.9 C -1924.6 2024.7 1922.6 2014.5 1917.4 2020.4 C -1915.5 2022.8 1912 2022.6 1910.9 2025.4 C -1911.5 2031.9 1910.9 2038.8 1911.4 2045.3 C -1911.1 2046.5 1910 2047.4 1910.4 2048.9 C -1915.1 2054.4 1920.4 2058.3 1925.1 2063.8 C -1920.8 2058.6 1914.9 2054.3 1910.7 2048.9 C -[0.4 0.4 0 0] vc -f -S -n -1934.7 2031.9 m -1934.6 2030.7 1934.9 2029.5 1934.4 2028.5 C -1934 2029.5 1934.3 2031.2 1934.2 2032.6 C -1933.8 2031.7 1934.9 2031.6 1934.7 2031.9 C -[0.92 0.92 0 0.67] vc -f -S -n -vmrs -1934.7 2019.4 m -1934.1 2015.3 1935.6 2010.9 1934.9 2007.9 C -1935.1 2007.8 1935.6 2008.1 1935.4 2007.6 C -1936.8 2008.6 1938.2 2007 1939.7 2006.2 C -1940.1 2004.3 1942.7 2005 1943.6 2003.8 C -1945.1 2000.3 1954 2000.8 1950 1996.6 C -1952.1 1993.3 1948.2 1989.2 1951.2 1985.6 C -1953 1981.4 1948.4 1982.3 1947.9 1979.8 C -1945.4 1979.6 1945.1 1975.5 1942.4 1975 C -1942.4 1972.3 1938 1973.6 1938.5 1970.4 C -1937.4 1969 1935.6 1970.1 1934.2 1970.2 C -1927.5 1974.5 1919.8 1976.8 1913.8 1982.2 C -1913.8 1990.4 1913.8 1998.7 1913.8 2006.9 C -1919.3 2012.7 1923.8 2016.2 1929.2 2022.3 C -1931.1 2021.6 1932.8 2018.9 1934.7 2019.4 C -[0 0 0 0] vc -f -0.4 w -2 J -2 M -S -n -2024.2 2038.1 m -2024.1 2029.3 2024.4 2021.7 2024.7 2014.4 C -2024.4 2013.6 2020.6 2013.4 2021.3 2011.2 C -2020.5 2010.3 2018.4 2010.6 2018.9 2008.6 C -2019 2008.8 2018.8 2009 2018.7 2009.1 C -2018.2 2006.7 2015.2 2007.9 2015.3 2005.5 C -2014.7 2004.8 2012.4 2005.1 2013.2 2003.6 C -2012.3 2004.2 2012.8 2002.4 2012.7 2002.6 C -2009.4 2003.3 2011.2 1998.6 2008.4 1999.2 C -2007 1999.1 2006.1 1999.4 2005.7 2000.4 C -2006.9 1998.5 2007.7 2000.5 2009.3 2000.2 C -2009.2 2003.7 2012.4 2002.1 2012.9 2005.2 C -2015.9 2005.6 2015.2 2008.6 2017.7 2008.8 C -2018.4 2009.6 2018.3 2011.4 2019.6 2011 C -2021.1 2011.7 2021.4 2014.8 2023.7 2015.1 C -2023.7 2023.5 2023.9 2031.6 2023.5 2040.5 C -2021.8 2041.7 2020.7 2043.6 2018.4 2043.9 C -2020.8 2042.7 2025.5 2041.8 2024.2 2038.1 C -[0 0.87 0.91 0.83] vc -f -S -n -2023.5 2040 m -2023.5 2031.1 2023.5 2023.4 2023.5 2015.1 C -2020.2 2015 2021.8 2010.3 2018.4 2011 C -2018.6 2007.5 2014.7 2009.3 2014.8 2006.4 C -2011.8 2006.3 2012.2 2002.3 2009.8 2002.4 C -2009.7 2001.5 2009.2 2000.1 2008.4 2000.2 C -2008.7 2000.9 2009.7 2001.2 2009.3 2002.4 C -2008.4 2004.2 2007.5 2003.1 2007.9 2005.5 C -2007.9 2010.8 2007.7 2018.7 2008.1 2023.2 C -2009 2024.3 2007.3 2023.4 2007.9 2024 C -2007.7 2024.6 2007.3 2026.3 2008.6 2027.1 C -2009.7 2026.8 2010 2027.6 2010.5 2028 C -2010.5 2028.2 2010.5 2029.1 2010.5 2028.5 C -2011.5 2028 2010.5 2030 2011.5 2030 C -2014.2 2029.7 2012.9 2032.2 2014.8 2032.6 C -2015.1 2033.6 2015.3 2033 2016 2033.3 C -2017 2033.9 2016.6 2035.4 2017.2 2036.2 C -2018.7 2036.4 2019.2 2039 2021.3 2038.4 C -2021.6 2035.4 2019.7 2029.5 2021.1 2027.3 C -2020.9 2023.5 2021.5 2018.5 2020.6 2016 C -2020.9 2013.9 2021.5 2015.4 2022.3 2014.4 C -2022.2 2015.1 2023.3 2014.8 2023.2 2015.6 C -2022.7 2019.8 2023.3 2024.3 2022.8 2028.5 C -2022.3 2028.2 2022.6 2027.6 2022.5 2027.1 C -2022.5 2027.8 2022.5 2029.2 2022.5 2029.2 C -2022.6 2029.2 2022.7 2029.1 2022.8 2029 C -2023.9 2032.8 2022.6 2037 2023 2040.8 C -2022.3 2041.2 2021.6 2041.5 2021.1 2042.2 C -2022 2041.2 2022.9 2041.4 2023.5 2040 C -[0 1 1 0.23] vc -f -S -n -2009.1 1997.8 m -2003.8 1997.7 2000.1 2002.4 1995.4 2003.1 C -1995 1999.5 1995.2 1995 1995.2 1992 C -1995.2 1995.8 1995 1999.7 1995.4 2003.3 C -2000.3 2002.2 2003.8 1997.9 2009.1 1997.8 C -2012.3 2001.2 2015.6 2004.8 2018.7 2008.1 C -2021.6 2011.2 2027.5 2013.9 2025.9 2019.9 C -2026.1 2017.9 2025.6 2016.2 2025.4 2014.4 C -2020.2 2008.4 2014 2003.6 2009.1 1997.8 C -[0.18 0.18 0 0.78] vc -f -S -n -2009.3 1997.8 m -2008.7 1997.4 2007.9 1997.6 2007.2 1997.6 C -2007.9 1997.6 2008.9 1997.4 2009.6 1997.8 C -2014.7 2003.6 2020.8 2008.8 2025.9 2014.8 C -2025.8 2017.7 2026.1 2014.8 2025.6 2014.1 C -2020.4 2008.8 2014.8 2003.3 2009.3 1997.8 C -[0.07 0.06 0 0.58] vc -f -S -n -2009.6 1997.6 m -2009 1997.1 2008.1 1997.4 2007.4 1997.3 C -2008.1 1997.4 2009 1997.1 2009.6 1997.6 C -2014.8 2003.7 2021.1 2008.3 2025.9 2014.4 C -2021.1 2008.3 2014.7 2003.5 2009.6 1997.6 C -[0.4 0.4 0 0] vc -f -S -n -2021.8 2011.5 m -2021.9 2012.2 2022.3 2013.5 2023.7 2013.6 C -2023.4 2012.7 2022.8 2011.8 2021.8 2011.5 C -[0 0.33 0.33 0.99] vc -f -S -n -2021.1 2042 m -2022.1 2041.1 2020.9 2040.2 2020.6 2039.6 C -2018.4 2039.5 2018.1 2036.9 2016.3 2036.4 C -2015.8 2035.5 2015.3 2033.8 2014.8 2033.6 C -2012.4 2033.8 2013 2030.4 2010.5 2030.2 C -2009.6 2028.9 2009.6 2028.3 2008.4 2028 C -2006.9 2026.7 2007.5 2024.3 2006 2023.2 C -2006.6 2023.2 2005.7 2023.3 2005.7 2023 C -2006.4 2022.5 2006.3 2021.1 2006.7 2020.6 C -2006.6 2015 2006.9 2009 2006.4 2003.8 C -2006.9 2002.5 2007.6 2001.1 2006.9 2000.7 C -2004.6 2003.6 2003 2002.9 2000.2 2004.3 C -1999.3 2005.8 1997.9 2006.3 1996.1 2006.7 C -1995.7 2008.9 1996 2011.1 1995.9 2012.9 C -1993.4 2015.1 1990.5 2016.2 1987.7 2017.7 C -1987.1 2019.3 1991.1 2019.4 1990.4 2021.3 C -1990.5 2021.5 1991.9 2022.3 1992 2023 C -1994.8 2024.4 1996.2 2027.5 1998.5 2030 C -2002.4 2033 2005.2 2037.2 2008.8 2041 C -2010.2 2041.3 2011.6 2042 2011 2043.9 C -2011.2 2044.8 2010.1 2045.3 2010.5 2046.3 C -2013.8 2044.8 2017.5 2043.4 2021.1 2042 C -[0 0.5 0.5 0.2] vc -f -S -n -2019.4 2008.8 m -2018.9 2009.2 2019.3 2009.9 2019.6 2010.3 C -2022.2 2011.5 2020.3 2009.1 2019.4 2008.8 C -[0 0.33 0.33 0.99] vc -f -S -n -2018 2007.4 m -2015.7 2006.7 2015.3 2003.6 2012.9 2002.8 C -2013.5 2003.7 2013.5 2005.1 2015.6 2005.2 C -2016.4 2006.1 2015.7 2007.7 2018 2007.4 C -f -S -n -vmrs -1993.5 2008.8 m -1993.4 2000 1993.7 1992.5 1994 1985.1 C -1993.7 1984.3 1989.9 1984.1 1990.6 1982 C -1989.8 1981.1 1987.7 1981.4 1988.2 1979.3 C -1988.3 1979.6 1988.1 1979.7 1988 1979.8 C -1987.5 1977.5 1984.5 1978.6 1984.6 1976.2 C -1983.9 1975.5 1981.7 1975.8 1982.4 1974.3 C -1981.6 1974.9 1982.1 1973.1 1982 1973.3 C -1979 1973.7 1980 1968.8 1976.9 1969.7 C -1975.9 1969.8 1975.3 1970.3 1975 1971.2 C -1976.2 1969.2 1977 1971.2 1978.6 1970.9 C -1978.5 1974.4 1981.7 1972.8 1982.2 1976 C -1985.2 1976.3 1984.5 1979.3 1987 1979.6 C -1987.7 1980.3 1987.5 1982.1 1988.9 1981.7 C -1990.4 1982.4 1990.7 1985.5 1993 1985.8 C -1992.9 1994.3 1993.2 2002.3 1992.8 2011.2 C -1991.1 2012.4 1990 2014.4 1987.7 2014.6 C -1990.1 2013.4 1994.7 2012.6 1993.5 2008.8 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1992.8 2010.8 m -1992.8 2001.8 1992.8 1994.1 1992.8 1985.8 C -1989.5 1985.7 1991.1 1981.1 1987.7 1981.7 C -1987.9 1978.2 1983.9 1980 1984.1 1977.2 C -1981.1 1977 1981.5 1973 1979.1 1973.1 C -1979 1972.2 1978.5 1970.9 1977.6 1970.9 C -1977.9 1971.6 1979 1971.9 1978.6 1973.1 C -1977.6 1974.9 1976.8 1973.9 1977.2 1976.2 C -1977.2 1981.5 1977 1989.4 1977.4 1994 C -1978.3 1995 1976.6 1994.1 1977.2 1994.7 C -1977 1995.3 1976.6 1997 1977.9 1997.8 C -1979 1997.5 1979.3 1998.3 1979.8 1998.8 C -1979.8 1998.9 1979.8 1999.8 1979.8 1999.2 C -1980.8 1998.7 1979.7 2000.7 1980.8 2000.7 C -1983.5 2000.4 1982.1 2003 1984.1 2003.3 C -1984.4 2004.3 1984.5 2003.7 1985.3 2004 C -1986.3 2004.6 1985.9 2006.1 1986.5 2006.9 C -1988 2007.1 1988.4 2009.7 1990.6 2009.1 C -1990.9 2006.1 1989 2000.2 1990.4 1998 C -1990.2 1994.3 1990.8 1989.2 1989.9 1986.8 C -1990.2 1984.7 1990.8 1986.2 1991.6 1985.1 C -1991.5 1985.9 1992.6 1985.5 1992.5 1986.3 C -1992 1990.5 1992.6 1995 1992 1999.2 C -1991.6 1998.9 1991.9 1998.3 1991.8 1997.8 C -1991.8 1998.5 1991.8 2000 1991.8 2000 C -1991.9 1999.9 1992 1999.8 1992 1999.7 C -1993.2 2003.5 1991.9 2007.7 1992.3 2011.5 C -1991.6 2012 1990.9 2012.2 1990.4 2012.9 C -1991.3 2011.9 1992.2 2012.1 1992.8 2010.8 C -[0 1 1 0.23] vc -f -S -n -1978.4 1968.5 m -1977 1969.2 1975.8 1968.2 1974.5 1969 C -1968.3 1973 1961.6 1976 1955.1 1979.1 C -1962 1975.9 1968.8 1972.5 1975.5 1968.8 C -1976.5 1968.8 1977.6 1968.8 1978.6 1968.8 C -1981.7 1972.1 1984.8 1975.7 1988 1978.8 C -1990.9 1981.9 1996.8 1984.6 1995.2 1990.6 C -1995.3 1988.6 1994.9 1986.9 1994.7 1985.1 C -1989.5 1979.1 1983.3 1974.3 1978.4 1968.5 C -[0.18 0.18 0 0.78] vc -f -S -n -1978.4 1968.3 m -1977.9 1968.7 1977.1 1968.5 1976.4 1968.5 C -1977.3 1968.8 1978.1 1967.9 1978.8 1968.5 C -1984 1974.3 1990.1 1979.5 1995.2 1985.6 C -1995.1 1988.4 1995.3 1985.6 1994.9 1984.8 C -1989.5 1979.4 1983.9 1973.8 1978.4 1968.3 C -[0.07 0.06 0 0.58] vc -f -S -n -1978.6 1968 m -1977.9 1968 1977.4 1968.6 1978.4 1968 C -1983.9 1973.9 1990.1 1979.1 1995.2 1985.1 C -1990.2 1979 1983.8 1974.1 1978.6 1968 C -[0.4 0.4 0 0] vc -f -S -n -1991.1 1982.2 m -1991.2 1982.9 1991.6 1984.2 1993 1984.4 C -1992.6 1983.5 1992.1 1982.5 1991.1 1982.2 C -[0 0.33 0.33 0.99] vc -f -S -n -1990.4 2012.7 m -1991.4 2011.8 1990.2 2010.9 1989.9 2010.3 C -1987.7 2010.2 1987.4 2007.6 1985.6 2007.2 C -1985.1 2006.2 1984.6 2004.5 1984.1 2004.3 C -1981.7 2004.5 1982.3 2001.2 1979.8 2000.9 C -1978.8 1999.6 1978.8 1999.1 1977.6 1998.8 C -1976.1 1997.4 1976.7 1995 1975.2 1994 C -1975.8 1994 1975 1994 1975 1993.7 C -1975.7 1993.2 1975.6 1991.8 1976 1991.3 C -1975.9 1985.7 1976.1 1979.7 1975.7 1974.5 C -1976.2 1973.3 1976.9 1971.8 1976.2 1971.4 C -1973.9 1974.3 1972.2 1973.6 1969.5 1975 C -1967.9 1977.5 1963.8 1977.1 1961.8 1980 C -1959 1980 1957.6 1983 1954.8 1982.9 C -1953.8 1984.2 1954.8 1985.7 1955.1 1987.2 C -1956.2 1989.5 1959.7 1990.1 1959.9 1991.8 C -1965.9 1998 1971.8 2005.2 1978.1 2011.7 C -1979.5 2012 1980.9 2012.7 1980.3 2014.6 C -1980.5 2015.6 1979.4 2016 1979.8 2017 C -1983 2015.6 1986.8 2014.1 1990.4 2012.7 C -[0 0.5 0.5 0.2] vc -f -S -n -1988.7 1979.6 m -1988.2 1979.9 1988.6 1980.6 1988.9 1981 C -1991.4 1982.2 1989.6 1979.9 1988.7 1979.6 C -[0 0.33 0.33 0.99] vc -f -S -n -1987.2 1978.1 m -1985 1977.5 1984.6 1974.3 1982.2 1973.6 C -1982.7 1974.5 1982.8 1975.8 1984.8 1976 C -1985.7 1976.9 1985 1978.4 1987.2 1978.1 C -f -S -n -1975.5 2084 m -1975.5 2082 1975.3 2080 1975.7 2078.2 C -1978.8 2079 1980.9 2085.5 1984.8 2083.5 C -1993 2078.7 2001.6 2075 2010 2070.8 C -2010.1 2064 2009.9 2057.2 2010.3 2050.6 C -2014.8 2046.2 2020.9 2045.7 2025.6 2042 C -2026.1 2035.1 2025.8 2028 2025.9 2021.1 C -2025.8 2027.8 2026.1 2034.6 2025.6 2041.2 C -2022.2 2044.9 2017.6 2046.8 2012.9 2048 C -2012.5 2049.5 2010.4 2049.4 2009.8 2051.1 C -2009.9 2057.6 2009.6 2064.2 2010 2070.5 C -2001.2 2075.4 1992 2079.1 1983.2 2084 C -1980.3 2082.3 1977.8 2079.2 1975.2 2077.5 C -1974.9 2079.9 1977.2 2084.6 1973.3 2085.2 C -1964.7 2088.6 1956.8 2093.7 1948.1 2097.2 C -1949 2097.3 1949.6 2096.9 1950.3 2096.7 C -1958.4 2091.9 1967.1 2088.2 1975.5 2084 C -[0.18 0.18 0 0.78] vc -f -S -n -vmrs -1948.6 2094.5 m -1950.2 2093.7 1951.8 2092.9 1953.4 2092.1 C -1951.8 2092.9 1950.2 2093.7 1948.6 2094.5 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1971.6 2082.3 m -1971.6 2081.9 1970.7 2081.1 1970.9 2081.3 C -1970.7 2081.6 1970.6 2081.6 1970.4 2081.3 C -1970.8 2080.1 1968.7 2081.7 1968.3 2080.8 C -1966.6 2080.9 1966.7 2078 1964.2 2078.2 C -1964.8 2075 1960.1 2075.8 1960.1 2072.9 C -1958 2072.3 1957.5 2069.3 1955.3 2069.3 C -1953.9 2070.9 1948.8 2067.8 1950 2072 C -1949 2074 1943.2 2070.6 1944 2074.8 C -1942.2 2076.6 1937.6 2073.9 1938 2078.2 C -1936.7 2078.6 1935 2078.6 1933.7 2078.2 C -1933.5 2080 1936.8 2080.7 1937.3 2082.8 C -1939.9 2083.5 1940.6 2086.4 1942.6 2088 C -1945.2 2089.2 1946 2091.3 1948.4 2093.6 C -1956 2089.5 1963.9 2086.1 1971.6 2082.3 C -[0 0.01 1 0] vc -f -S -n -1958.2 2089.7 m -1956.4 2090 1955.6 2091.3 1953.9 2091.9 C -1955.6 2091.9 1956.5 2089.7 1958.2 2089.7 C -[0 0.87 0.91 0.83] vc -f -S -n -1929.9 2080.4 m -1929.5 2077.3 1929.7 2073.9 1929.6 2070.8 C -1929.8 2074.1 1929.2 2077.8 1930.1 2080.8 C -1935.8 2085.9 1941.4 2091.3 1946.9 2096.9 C -1941.2 2091 1935.7 2086 1929.9 2080.4 C -[0.4 0.4 0 0] vc -f -S -n -1930.1 2080.4 m -1935.8 2086 1941.5 2090.7 1946.9 2096.7 C -1941.5 2090.9 1935.7 2085.8 1930.1 2080.4 C -[0.07 0.06 0 0.58] vc -f -S -n -1940.9 2087.1 m -1941.7 2088 1944.8 2090.6 1943.6 2089.2 C -1942.5 2089 1941.6 2087.7 1940.9 2087.1 C -[0 0.87 0.91 0.83] vc -f -S -n -1972.8 2082.8 m -1973 2075.3 1972.4 2066.9 1973.3 2059.5 C -1972.5 2058.9 1972.8 2057.3 1973.1 2056.4 C -1974.8 2055.2 1973.4 2055.5 1972.4 2055.4 C -1970.1 2053.2 1967.9 2050.9 1965.6 2048.7 C -1960.9 2049.9 1956.9 2052.7 1952.4 2054.7 C -1949.3 2052.5 1946.3 2049.5 1943.6 2046.8 C -1939.9 2047.7 1936.8 2050.1 1933.5 2051.8 C -1930.9 2054.9 1933.5 2056.2 1932.3 2059.7 C -1933.2 2059.7 1932.2 2060.5 1932.5 2060.2 C -1933.2 2062.5 1931.6 2064.6 1932.5 2067.4 C -1932.9 2069.7 1932.7 2072.2 1932.8 2074.6 C -1933.6 2070.6 1932.2 2066.3 1933 2062.6 C -1934.4 2058.2 1929.8 2053.5 1935.2 2051.1 C -1937.7 2049.7 1940.2 2048 1942.8 2046.8 C -1945.9 2049.2 1948.8 2052 1951.7 2054.7 C -1952.7 2054.7 1953.6 2054.6 1954.4 2054.2 C -1958.1 2052.5 1961.7 2049.3 1965.9 2049.2 C -1968.2 2052.8 1975.2 2055 1972.6 2060.9 C -1973.3 2062.4 1972.2 2065.2 1972.6 2067.6 C -1972.7 2072.6 1972.4 2077.7 1972.8 2082.5 C -1968.1 2084.9 1963.5 2087.5 1958.7 2089.5 C -1963.5 2087.4 1968.2 2085 1972.8 2082.8 C -f -S -n -1935.2 2081.1 m -1936.8 2083.4 1938.6 2084.6 1940.4 2086.6 C -1938.8 2084.4 1936.7 2083.4 1935.2 2081.1 C -f -S -n -1983.2 2081.3 m -1984.8 2080.5 1986.3 2079.7 1988 2078.9 C -1986.3 2079.7 1984.8 2080.5 1983.2 2081.3 C -f -S -n -2006.2 2069.1 m -2006.2 2068.7 2005.2 2067.9 2005.5 2068.1 C -2005.3 2068.4 2005.2 2068.4 2005 2068.1 C -2005.4 2066.9 2003.3 2068.5 2002.8 2067.6 C -2001.2 2067.7 2001.2 2064.8 1998.8 2065 C -1999.4 2061.8 1994.7 2062.6 1994.7 2059.7 C -1992.4 2059.5 1992.4 2055.8 1990.1 2056.8 C -1985.9 2059.5 1981.1 2061 1976.9 2063.8 C -1977.2 2067.6 1974.9 2074.2 1978.8 2075.8 C -1979.6 2077.8 1981.7 2078.4 1982.9 2080.4 C -1990.6 2076.3 1998.5 2072.9 2006.2 2069.1 C -[0 0.01 1 0] vc -f -S -n -vmrs -1992.8 2076.5 m -1991 2076.8 1990.2 2078.1 1988.4 2078.7 C -1990.2 2078.7 1991 2076.5 1992.8 2076.5 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1975.5 2073.4 m -1976.1 2069.7 1973.9 2064.6 1977.4 2062.4 C -1973.9 2064.5 1976.1 2069.9 1975.5 2073.6 C -1976 2074.8 1979.3 2077.4 1978.1 2076 C -1977 2075.7 1975.8 2074.5 1975.5 2073.4 C -f -S -n -2007.4 2069.6 m -2007.6 2062.1 2007 2053.7 2007.9 2046.3 C -2007.1 2045.7 2007.3 2044.1 2007.6 2043.2 C -2009.4 2042 2007.9 2042.3 2006.9 2042.2 C -2002.2 2037.4 1996.7 2032.4 1992.5 2027.3 C -1992 2027.3 1991.6 2027.3 1991.1 2027.3 C -1991.4 2035.6 1991.4 2045.6 1991.1 2054.4 C -1990.5 2055.5 1988.4 2056.6 1990.6 2055.4 C -1991.6 2055.4 1991.6 2054.1 1991.6 2053.2 C -1990.8 2044.7 1991.9 2035.4 1991.6 2027.6 C -1991.8 2027.6 1992 2027.6 1992.3 2027.6 C -1997 2032.8 2002.5 2037.7 2007.2 2042.9 C -2007.3 2044.8 2006.7 2047.4 2007.6 2048.4 C -2006.9 2055.1 2007.1 2062.5 2007.4 2069.3 C -2002.7 2071.7 1998.1 2074.3 1993.2 2076.3 C -1998 2074.2 2002.7 2071.8 2007.4 2069.6 C -f -S -n -2006.7 2069.1 m -2006.3 2068.6 2005.9 2067.7 2005.7 2066.9 C -2005.7 2059.7 2005.9 2051.4 2005.5 2045.1 C -2004.9 2045.3 2004.7 2044.5 2004.3 2045.3 C -2005.1 2045.3 2004.2 2045.8 2004.8 2046 C -2004.8 2052.2 2004.8 2059.2 2004.8 2064.5 C -2005.7 2065.7 2005.1 2065.7 2005 2066.7 C -2003.8 2067 2002.7 2067.2 2001.9 2066.4 C -2001.3 2064.6 1998 2063.1 1998 2061.9 C -1996.1 2062.3 1996.6 2058.3 1994.2 2058.8 C -1992.6 2057.7 1992.7 2054.8 1989.9 2056.6 C -1985.6 2059.3 1980.9 2060.8 1976.7 2063.6 C -1976 2066.9 1976 2071.2 1976.7 2074.6 C -1977.6 2070.8 1973.1 2062.1 1980.5 2061.2 C -1984.3 2060.3 1987.5 2058.2 1990.8 2056.4 C -1991.7 2056.8 1992.9 2057.2 1993.5 2059.2 C -1994.3 2058.6 1994.4 2060.6 1994.7 2059.2 C -1995.3 2062.7 1999.2 2061.4 1998.8 2064.8 C -2001.8 2065.4 2002.5 2068.4 2005.2 2067.4 C -2004.9 2067.9 2006 2068 2006.4 2069.1 C -2001.8 2071.1 1997.4 2073.9 1992.8 2075.8 C -1997.5 2073.8 2002 2071.2 2006.7 2069.1 C -[0 0.2 1 0] vc -f -S -n -1988.7 2056.6 m -1985.1 2058.7 1981.1 2060.1 1977.6 2061.9 C -1981.3 2060.5 1985.6 2058.1 1988.7 2056.6 C -[0 0.87 0.91 0.83] vc -f -S -n -1977.9 2059.5 m -1975.7 2064.5 1973.7 2054.7 1975.2 2060.9 C -1976 2060.6 1977.6 2059.7 1977.9 2059.5 C -f -S -n -1989.6 2051.3 m -1990.1 2042.3 1989.8 2036.6 1989.9 2028 C -1989.8 2027 1990.8 2028.3 1990.1 2027.3 C -1988.9 2026.7 1986.7 2026.9 1986.8 2024.7 C -1987.4 2023 1985.9 2024.6 1985.1 2023.7 C -1984.1 2021.4 1982.5 2020.5 1980.3 2020.6 C -1979.9 2020.8 1979.5 2021.1 1979.3 2021.6 C -1979.7 2025.8 1978.4 2033 1979.6 2038.1 C -1983.7 2042.9 1968.8 2044.6 1978.8 2042.7 C -1979.3 2042.3 1979.6 2041.9 1980 2041.5 C -1980 2034.8 1980 2027 1980 2021.6 C -1981.3 2020.5 1981.7 2021.5 1982.9 2021.8 C -1983.6 2024.7 1986.1 2023.8 1986.8 2026.4 C -1987.1 2027.7 1988.6 2027.1 1989.2 2028.3 C -1989.1 2036.7 1989.3 2044.8 1988.9 2053.7 C -1987.2 2054.9 1986.2 2056.8 1983.9 2057.1 C -1986.3 2055.9 1990.9 2055 1989.6 2051.3 C -f -S -n -1971.6 2078.9 m -1971.4 2070.5 1972.1 2062.2 1971.6 2055.9 C -1969.9 2053.7 1967.6 2051.7 1965.6 2049.6 C -1961.4 2050.4 1957.6 2053.6 1953.4 2055.2 C -1949.8 2055.6 1948.2 2051.2 1945.5 2049.6 C -1945.1 2048.8 1944.5 2047.9 1943.6 2047.5 C -1940.1 2047.8 1937.3 2051 1934 2052.3 C -1933.7 2052.6 1933.7 2053 1933.2 2053.2 C -1933.7 2060.8 1933.4 2067.2 1933.5 2074.6 C -1933.8 2068.1 1934 2060.9 1933.2 2054 C -1935.3 2050.9 1939.3 2049.6 1942.4 2047.5 C -1942.8 2047.5 1943.4 2047.4 1943.8 2047.7 C -1947.1 2050.2 1950.3 2057.9 1955.3 2054.4 C -1955.4 2054.4 1955.5 2054.3 1955.6 2054.2 C -1955.9 2057.6 1956.1 2061.8 1955.3 2064.8 C -1955.4 2064.3 1955.1 2063.8 1955.6 2063.6 C -1956 2066.6 1955.3 2068.7 1958.7 2069.8 C -1959.2 2071.7 1961.4 2071.7 1962 2074.1 C -1964.4 2074.2 1964 2077.7 1967.3 2078.4 C -1967 2079.7 1968.1 2079.9 1969 2080.1 C -1971.1 2079.9 1970 2079.2 1970.4 2078 C -1969.5 2077.2 1970.3 2075.9 1969.7 2075.1 C -1970.1 2069.8 1970.1 2063.6 1969.7 2058.8 C -1969.2 2058.5 1970 2058.1 1970.2 2057.8 C -1970.4 2058.3 1971.2 2057.7 1971.4 2058.3 C -1971.5 2065.3 1971.2 2073.6 1971.6 2081.1 C -1974.1 2081.4 1969.8 2084.3 1972.4 2082.5 C -1971.9 2081.4 1971.6 2080.2 1971.6 2078.9 C -[0 0.4 1 0] vc -f -S -n -1952.4 2052 m -1954.1 2051.3 1955.6 2050.4 1957.2 2049.6 C -1955.6 2050.4 1954.1 2051.3 1952.4 2052 C -[0 0.87 0.91 0.83] vc -f -S -n -1975.5 2039.8 m -1975.5 2039.4 1974.5 2038.7 1974.8 2038.8 C -1974.6 2039.1 1974.5 2039.1 1974.3 2038.8 C -1974.6 2037.6 1972.5 2039.3 1972.1 2038.4 C -1970.4 2038.4 1970.5 2035.5 1968 2035.7 C -1968.6 2032.5 1964 2033.3 1964 2030.4 C -1961.9 2029.8 1961.4 2026.8 1959.2 2026.8 C -1957.7 2028.5 1952.6 2025.3 1953.9 2029.5 C -1952.9 2031.5 1947 2028.2 1947.9 2032.4 C -1946 2034.2 1941.5 2031.5 1941.9 2035.7 C -1940.6 2036.1 1938.9 2036.1 1937.6 2035.7 C -1937.3 2037.5 1940.7 2038.2 1941.2 2040.3 C -1943.7 2041.1 1944.4 2043.9 1946.4 2045.6 C -1949.1 2046.7 1949.9 2048.8 1952.2 2051.1 C -1959.9 2047.1 1967.7 2043.6 1975.5 2039.8 C -[0 0.01 1 0] vc -f -S -n -vmrs -1962 2047.2 m -1960.2 2047.5 1959.5 2048.9 1957.7 2049.4 C -1959.5 2049.5 1960.3 2047.2 1962 2047.2 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -2012.4 2046.3 m -2010.3 2051.3 2008.3 2041.5 2009.8 2047.7 C -2010.5 2047.4 2012.2 2046.5 2012.4 2046.3 C -f -S -n -1944.8 2044.6 m -1945.5 2045.6 1948.6 2048.1 1947.4 2046.8 C -1946.3 2046.5 1945.5 2045.2 1944.8 2044.6 C -f -S -n -1987.2 2054.9 m -1983.7 2057.3 1979.6 2058 1976 2060.2 C -1974.7 2058.2 1977.2 2055.8 1974.3 2054.9 C -1973.1 2052 1970.4 2050.2 1968 2048 C -1968 2047.7 1968 2047.4 1968.3 2047.2 C -1969.5 2046.1 1983 2040.8 1972.4 2044.8 C -1971.2 2046.6 1967.9 2046 1968 2048.2 C -1970.5 2050.7 1973.8 2052.6 1974.3 2055.6 C -1975.1 2055 1975.7 2056.7 1975.7 2057.1 C -1975.7 2058.2 1974.8 2059.3 1975.5 2060.4 C -1979.3 2058.2 1983.9 2057.7 1987.2 2054.9 C -[0.18 0.18 0 0.78] vc -f -S -n -1967.8 2047.5 m -1968.5 2047 1969.1 2046.5 1969.7 2046 C -1969.1 2046.5 1968.5 2047 1967.8 2047.5 C -[0 0.87 0.91 0.83] vc -f -S -n -1976.7 2040.3 m -1976.9 2032.8 1976.3 2024.4 1977.2 2017 C -1976.4 2016.5 1976.6 2014.8 1976.9 2013.9 C -1978.7 2012.7 1977.2 2013 1976.2 2012.9 C -1971.5 2008.1 1965.9 2003.1 1961.8 1998 C -1960.9 1998 1960.1 1998 1959.2 1998 C -1951.5 2001.1 1944.3 2005.5 1937.1 2009.6 C -1935 2012.9 1937 2013.6 1936.1 2017.2 C -1937.1 2017.2 1936 2018 1936.4 2017.7 C -1937 2020.1 1935.5 2022.1 1936.4 2024.9 C -1936.8 2027.2 1936.5 2029.7 1936.6 2032.1 C -1937.4 2028.2 1936 2023.8 1936.8 2020.1 C -1938.3 2015.7 1933.6 2011 1939 2008.6 C -1945.9 2004.5 1953.1 2000.3 1960.6 1998.3 C -1960.9 1998.3 1961.3 1998.3 1961.6 1998.3 C -1966.2 2003.5 1971.8 2008.4 1976.4 2013.6 C -1976.6 2015.5 1976 2018.1 1976.9 2019.2 C -1976.1 2025.8 1976.4 2033.2 1976.7 2040 C -1971.9 2042.4 1967.4 2045 1962.5 2047 C -1967.3 2044.9 1972 2042.6 1976.7 2040.3 C -f -S -n -1939 2038.6 m -1940.6 2040.9 1942.5 2042.1 1944.3 2044.1 C -1942.7 2041.9 1940.6 2040.9 1939 2038.6 C -f -S -n -2006.2 2065.7 m -2006 2057.3 2006.7 2049 2006.2 2042.7 C -2002.1 2038.4 1997.7 2033.4 1993 2030 C -1992.9 2029.3 1992.5 2028.6 1992 2028.3 C -1992.1 2036.6 1991.9 2046.2 1992.3 2054.9 C -1990.8 2056.2 1989 2056.7 1987.5 2058 C -1988.7 2057.7 1990.7 2054.4 1993 2056.4 C -1993.4 2058.8 1996 2058.2 1996.6 2060.9 C -1999 2061 1998.5 2064.5 2001.9 2065.2 C -2001.5 2066.5 2002.7 2066.7 2003.6 2066.9 C -2005.7 2066.7 2004.6 2066 2005 2064.8 C -2004 2064 2004.8 2062.7 2004.3 2061.9 C -2004.6 2056.6 2004.6 2050.4 2004.3 2045.6 C -2003.7 2045.3 2004.6 2044.9 2004.8 2044.6 C -2005 2045.1 2005.7 2044.5 2006 2045.1 C -2006 2052.1 2005.8 2060.4 2006.2 2067.9 C -2008.7 2068.2 2004.4 2071.1 2006.9 2069.3 C -2006.4 2068.2 2006.2 2067 2006.2 2065.7 C -[0 0.4 1 0] vc -f -S -n -2021.8 2041.7 m -2018.3 2044.1 2014.1 2044.8 2010.5 2047 C -2009.3 2045 2011.7 2042.6 2008.8 2041.7 C -2004.3 2035.1 1997.6 2030.9 1993 2024.4 C -1992.1 2024 1991.5 2024.3 1990.8 2024 C -1993.2 2023.9 1995.3 2027.1 1996.8 2029 C -2000.4 2032.6 2004.9 2036.9 2008.4 2040.8 C -2008.2 2043.1 2011.4 2042.8 2009.8 2045.8 C -2009.8 2046.3 2009.7 2046.9 2010 2047.2 C -2013.8 2045 2018.5 2044.5 2021.8 2041.7 C -[0.18 0.18 0 0.78] vc -f -S -n -2001.6 2034 m -2000.7 2033.1 1999.9 2032.3 1999 2031.4 C -1999.9 2032.3 2000.7 2033.1 2001.6 2034 C -[0 0.87 0.91 0.83] vc -f -S -n -vmrs -1989.4 2024.4 m -1989.5 2025.4 1988.6 2024.3 1988.9 2024.7 C -1990.5 2025.8 1990.7 2024.2 1992.8 2024.9 C -1993.8 2025.9 1995 2027.1 1995.9 2028 C -1994.3 2026 1991.9 2023.4 1989.4 2024.4 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1984.8 2019.9 m -1984.6 2018.6 1986.3 2017.2 1987.7 2016.8 C -1987.2 2017.5 1982.9 2017.9 1984.4 2020.6 C -1984.1 2019.9 1984.9 2020 1984.8 2019.9 C -f -S -n -1981.7 2017 m -1979.6 2022 1977.6 2012.3 1979.1 2018.4 C -1979.8 2018.1 1981.5 2017.2 1981.7 2017 C -f -S -n -1884.3 2019.2 m -1884.7 2010.5 1884.5 2000.6 1884.5 1991.8 C -1886.6 1989.3 1889.9 1988.9 1892.4 1987 C -1890.8 1988.7 1886 1989.1 1884.3 1992.3 C -1884.7 2001 1884.5 2011.3 1884.5 2019.9 C -1891 2025.1 1895.7 2031.5 1902 2036.9 C -1896.1 2031 1890 2024.9 1884.3 2019.2 C -[0.07 0.06 0 0.58] vc -f -S -n -1884 2019.4 m -1884.5 2010.6 1884.2 2000.4 1884.3 1991.8 C -1884.8 1990.4 1887.8 1989 1884.8 1990.8 C -1884.3 1991.3 1884.3 1992 1884 1992.5 C -1884.5 2001.2 1884.2 2011.1 1884.3 2019.9 C -1887.9 2023.1 1891.1 2026.4 1894.4 2030 C -1891.7 2026.1 1887.1 2022.9 1884 2019.4 C -[0.4 0.4 0 0] vc -f -S -n -1885 2011.7 m -1885 2006.9 1885 2001.9 1885 1997.1 C -1885 2001.9 1885 2006.9 1885 2011.7 C -[0 0.87 0.91 0.83] vc -f -S -n -1975.5 2036.4 m -1975.2 2028 1976 2019.7 1975.5 2013.4 C -1971.1 2008.5 1965.6 2003.6 1961.6 1999 C -1958.8 1998 1956 2000 1953.6 2001.2 C -1948.2 2004.7 1941.9 2006.5 1937.1 2010.8 C -1937.5 2018.3 1937.3 2024.7 1937.3 2032.1 C -1937.6 2025.6 1937.9 2018.4 1937.1 2011.5 C -1937.3 2011 1937.6 2010.5 1937.8 2010 C -1944.6 2005.7 1951.9 2002.3 1959.2 1999 C -1960.1 1998.5 1960.1 1999.8 1960.4 2000.4 C -1959.7 2006.9 1959.7 2014.2 1959.4 2021.1 C -1959 2021.1 1959.2 2021.9 1959.2 2022.3 C -1959.2 2021.9 1959 2021.3 1959.4 2021.1 C -1959.8 2024.1 1959.2 2026.2 1962.5 2027.3 C -1963 2029.2 1965.3 2029.2 1965.9 2031.6 C -1968.3 2031.8 1967.8 2035.2 1971.2 2036 C -1970.8 2037.2 1971.9 2037.5 1972.8 2037.6 C -1974.9 2037.4 1973.9 2036.7 1974.3 2035.5 C -1973.3 2034.7 1974.1 2033.4 1973.6 2032.6 C -1973.9 2027.3 1973.9 2021.1 1973.6 2016.3 C -1973 2016 1973.9 2015.6 1974 2015.3 C -1974.3 2015.9 1975 2015.3 1975.2 2015.8 C -1975.3 2022.8 1975.1 2031.2 1975.5 2038.6 C -1977.9 2039 1973.7 2041.8 1976.2 2040 C -1975.7 2039 1975.5 2037.8 1975.5 2036.4 C -[0 0.4 1 0] vc -f -S -n -1991.1 2012.4 m -1987.5 2014.8 1983.4 2015.6 1979.8 2017.7 C -1978.5 2015.7 1981 2013.3 1978.1 2012.4 C -1973.6 2005.8 1966.8 2001.6 1962.3 1995.2 C -1961.4 1994.7 1960.8 1995 1960.1 1994.7 C -1962.5 1994.6 1964.6 1997.8 1966.1 1999.7 C -1969.7 2003.3 1974.2 2007.6 1977.6 2011.5 C -1977.5 2013.8 1980.6 2013.5 1979.1 2016.5 C -1979.1 2017 1979 2017.6 1979.3 2018 C -1983.1 2015.7 1987.8 2015.2 1991.1 2012.4 C -[0.18 0.18 0 0.78] vc -f -S -n -1970.9 2004.8 m -1970 2003.9 1969.2 2003 1968.3 2002.1 C -1969.2 2003 1970 2003.9 1970.9 2004.8 C -[0 0.87 0.91 0.83] vc -f -S -n -1887.9 1994.9 m -1888.5 1992.3 1891.4 1992.2 1893.2 1990.8 C -1898.4 1987.5 1904 1984.8 1909.5 1982.2 C -1909.7 1982.7 1910.3 1982.1 1910.4 1982.7 C -1909.5 1990.5 1910.1 1996.4 1910 2004.5 C -1909.1 2003.4 1909.7 2005.8 1909.5 2006.4 C -1910.4 2006 1909.7 2008 1910.2 2007.9 C -1911.3 2010.6 1912.5 2012.6 1915.7 2013.4 C -1915.8 2013.7 1915.5 2014.4 1916 2014.4 C -1916.3 2015 1915.4 2016 1915.2 2016 C -1916.1 2015.5 1916.5 2014.5 1916 2013.6 C -1913.4 2013.3 1913.1 2010.5 1910.9 2009.8 C -1910.7 2008.8 1910.4 2007.9 1910.2 2006.9 C -1911.1 1998.8 1909.4 1990.7 1910.7 1982.4 C -1910 1982.1 1908.9 1982.1 1908.3 1982.4 C -1901.9 1986.1 1895 1988.7 1888.8 1993 C -1888 1993.4 1888.4 1994.3 1887.6 1994.7 C -1888.1 2001.3 1887.8 2008.6 1887.9 2015.1 C -1887.3 2017.5 1887.9 2015.4 1888.4 2014.4 C -1887.8 2008 1888.4 2001.3 1887.9 1994.9 C -[0.07 0.06 0 0.58] vc -f -S -n -vmrs -1887.9 2018.4 m -1887.5 2016.9 1888.5 2016 1888.8 2014.8 C -1890.1 2014.8 1891.1 2016.6 1892.4 2015.3 C -1892.4 2014.4 1893.8 2012.9 1894.4 2012.4 C -1895.9 2012.4 1896.6 2013.9 1897.7 2012.7 C -1898.4 2011.7 1898.6 2010.4 1899.6 2009.8 C -1901.7 2009.9 1902.9 2010.4 1904 2009.1 C -1904.3 2007.4 1904 2007.6 1904.9 2007.2 C -1906.2 2007 1907.6 2006.5 1908.8 2006.7 C -1910.6 2008.2 1909.8 2011.5 1912.6 2012 C -1912.4 2013 1913.8 2012.7 1914 2013.2 C -1911.5 2011.1 1909.1 2007.9 1909.2 2004.3 C -1909.5 2003.5 1909.9 2004.9 1909.7 2004.3 C -1909.9 1996.2 1909.3 1990.5 1910.2 1982.7 C -1909.5 1982.6 1909.5 1982.6 1908.8 1982.7 C -1903.1 1985.7 1897 1987.9 1891.7 1992 C -1890.5 1993 1888.2 1992.9 1888.1 1994.9 C -1888.7 2001.4 1888.1 2008.4 1888.6 2014.8 C -1888.3 2016 1887.2 2016.9 1887.6 2018.4 C -1892.3 2023.9 1897.6 2027.9 1902.3 2033.3 C -1898 2028.2 1892.1 2023.8 1887.9 2018.4 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1910.9 1995.2 m -1910.4 1999.8 1911 2003.3 1910.9 2008.1 C -1910.9 2003.8 1910.9 1999.2 1910.9 1995.2 C -[0.18 0.18 0 0.78] vc -f -S -n -1911.2 2004.3 m -1911.2 2001.9 1911.2 1999.7 1911.2 1997.3 C -1911.2 1999.7 1911.2 2001.9 1911.2 2004.3 C -[0 0.87 0.91 0.83] vc -f -S -n -1958.7 1995.2 m -1959 1995.6 1956.2 1995 1956.5 1996.8 C -1955.8 1997.6 1954.2 1998.5 1953.6 1997.3 C -1953.6 1990.8 1954.9 1989.6 1953.4 1983.9 C -1953.4 1983.3 1953.3 1982.1 1954.4 1982 C -1955.5 1982.6 1956.5 1981.3 1957.5 1981 C -1956.3 1981.8 1954.7 1982.6 1953.9 1981.5 C -1951.4 1983 1954.7 1988.8 1952.9 1990.6 C -1953.8 1990.6 1953.2 1992.7 1953.4 1993.7 C -1953.8 1994.5 1952.3 1996.1 1953.2 1997.8 C -1956.3 1999.4 1957.5 1994 1959.9 1995.6 C -1962 1994.4 1963.7 1997.7 1965.2 1998.8 C -1963.5 1996.7 1961.2 1994.1 1958.7 1995.2 C -f -S -n -1945 2000.7 m -1945.4 1998.7 1945.4 1997.9 1945 1995.9 C -1944.5 1995.3 1944.2 1992.6 1945.7 1993.2 C -1946 1992.2 1948.7 1992.5 1948.4 1990.6 C -1947.5 1990.3 1948.1 1988.7 1947.9 1988.2 C -1948.9 1987.8 1950.5 1986.8 1950.5 1984.6 C -1951.5 1980.9 1946.7 1983 1947.2 1979.8 C -1944.5 1979.9 1945.2 1976.6 1943.1 1976.7 C -1941.8 1975.7 1942.1 1972.7 1939.2 1973.8 C -1938.2 1974.6 1939.3 1971.6 1938.3 1970.9 C -1938.8 1969.2 1933.4 1970.3 1937.3 1970 C -1939.4 1971.2 1937.2 1973 1937.6 1974.3 C -1937.2 1976.3 1937.1 1981.2 1937.8 1984.1 C -1938.8 1982.3 1937.9 1976.6 1938.5 1973.1 C -1938.9 1975 1938.5 1976.4 1939.7 1977.2 C -1939.5 1983.5 1938.9 1991.3 1940.2 1997.3 C -1939.4 1999.1 1938.6 1997.1 1937.8 1997.1 C -1937.4 1996.7 1937.6 1996.1 1937.6 1995.6 C -1936.5 1998.5 1940.1 1998.4 1940.9 2000.7 C -1942.1 2000.4 1943.2 2001.3 1943.1 2002.4 C -1943.6 2003.1 1941.1 2004.6 1942.8 2003.8 C -1943.9 2002.5 1942.6 2000.6 1945 2000.7 C -[0.65 0.65 0 0.42] vc -f -S -n -1914.5 2006.4 m -1914.1 2004.9 1915.2 2004 1915.5 2002.8 C -1916.7 2002.8 1917.8 2004.6 1919.1 2003.3 C -1919 2002.4 1920.4 2000.9 1921 2000.4 C -1922.5 2000.4 1923.2 2001.9 1924.4 2000.7 C -1925 1999.7 1925.3 1998.4 1926.3 1997.8 C -1928.4 1997.9 1929.5 1998.4 1930.6 1997.1 C -1930.9 1995.4 1930.7 1995.6 1931.6 1995.2 C -1932.8 1995 1934.3 1994.5 1935.4 1994.7 C -1936.1 1995.8 1936.9 1996.2 1936.6 1997.8 C -1938.9 1999.4 1939.7 2001.3 1942.4 2002.4 C -1942.4 2002.5 1942.2 2003 1942.6 2002.8 C -1942.9 2000.4 1939.2 2001.8 1939.2 1999.7 C -1936.2 1998.6 1937 1995.3 1935.9 1993.5 C -1937.1 1986.5 1935.2 1977.9 1937.6 1971.2 C -1937.6 1970.3 1936.6 1971 1936.4 1970.4 C -1930.2 1973.4 1924 1976 1918.4 1980 C -1917.2 1981 1914.9 1980.9 1914.8 1982.9 C -1915.3 1989.4 1914.7 1996.4 1915.2 2002.8 C -1914.9 2004 1913.9 2004.9 1914.3 2006.4 C -1919 2011.9 1924.2 2015.9 1928.9 2021.3 C -1924.6 2016.2 1918.7 2011.8 1914.5 2006.4 C -[0.4 0.4 0 0] vc -f -S -n -1914.5 1982.9 m -1915.1 1980.3 1918 1980.2 1919.8 1978.8 C -1925 1975.5 1930.6 1972.8 1936.1 1970.2 C -1939.4 1970.6 1936.1 1974.2 1936.6 1976.4 C -1936.5 1981.9 1936.8 1987.5 1936.4 1992.8 C -1935.9 1992.8 1936.2 1993.5 1936.1 1994 C -1937.1 1993.6 1936.2 1995.9 1936.8 1995.9 C -1937 1998 1939.5 1999.7 1940.4 2000.7 C -1940.1 1998.6 1935 1997.2 1937.6 1993.7 C -1938.3 1985.7 1935.9 1976.8 1937.8 1970.7 C -1936.9 1969.8 1935.4 1970.3 1934.4 1970.7 C -1928.3 1974.4 1921.4 1976.7 1915.5 1981 C -1914.6 1981.4 1915.1 1982.3 1914.3 1982.7 C -1914.7 1989.3 1914.5 1996.6 1914.5 2003.1 C -1913.9 2005.5 1914.5 2003.4 1915 2002.4 C -1914.5 1996 1915.1 1989.3 1914.5 1982.9 C -[0.07 0.06 0 0.58] vc -f -S -n -1939.2 1994.9 m -1939.3 1995 1939.4 1995.1 1939.5 1995.2 C -1939.1 1989 1939.3 1981.6 1939 1976.7 C -1938.6 1976.3 1938.6 1974.6 1938.5 1973.3 C -1938.7 1976.1 1938.1 1979.4 1939 1981.7 C -1937.3 1986 1937.7 1991.6 1938 1996.4 C -1937.3 1994.3 1939.6 1996.2 1939.2 1994.9 C -[0.18 0.18 0 0.78] vc -f -S -n -1938.3 1988.4 m -1938.5 1990.5 1937.9 1994.1 1938.8 1994.7 C -1937.9 1992.6 1939 1990.6 1938.3 1988.4 C -[0 0.87 0.91 0.83] vc -f -S -n -1938.8 1985.8 m -1938.5 1985.9 1938.4 1985.7 1938.3 1985.6 C -1938.4 1986.2 1938 1989.5 1938.8 1987.2 C -1938.8 1986.8 1938.8 1986.3 1938.8 1985.8 C -f -S -n -vmrs -1972.8 2062.1 m -1971.9 2061 1972.5 2059.4 1972.4 2058 C -1972.2 2063.8 1971.9 2073.7 1972.4 2081.3 C -1972.5 2074.9 1971.9 2067.9 1972.8 2062.1 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -S -n -1940.2 2071.7 m -1941.3 2072 1943.1 2072.3 1944 2071.5 C -1943.6 2069.9 1945.2 2069.1 1946 2068.8 C -1950 2071.1 1948.7 2065.9 1951.7 2066.2 C -1953.5 2063.9 1956.9 2069.4 1955.6 2063.8 C -1955.5 2064.2 1955.7 2064.8 1955.3 2065 C -1954.3 2063.7 1956.2 2063.6 1955.6 2062.1 C -1954.5 2060 1958.3 2050.3 1952.2 2055.6 C -1949.1 2053.8 1946 2051 1943.8 2048 C -1940.3 2048 1937.5 2051.3 1934.2 2052.5 C -1933.1 2054.6 1934.4 2057.3 1934 2060 C -1934 2065.1 1934 2069.7 1934 2074.6 C -1934.4 2069 1934.1 2061.5 1934.2 2054.9 C -1934.6 2054.5 1935.3 2054.7 1935.9 2054.7 C -1937 2055.3 1935.9 2056.1 1935.9 2056.8 C -1936.5 2063 1935.6 2070.5 1935.9 2074.6 C -1936.7 2074.4 1937.3 2075.2 1938 2074.6 C -1937.9 2073.6 1939.1 2072.1 1940.2 2071.7 C -[0 0.2 1 0] vc -f -S -n -1933.2 2074.1 m -1933.2 2071.5 1933.2 2069 1933.2 2066.4 C -1933.2 2069 1933.2 2071.5 1933.2 2074.1 C -[0 1 1 0.36] vc -f -S -n -2007.4 2048.9 m -2006.5 2047.8 2007.1 2046.2 2006.9 2044.8 C -2006.7 2050.6 2006.5 2060.5 2006.9 2068.1 C -2007.1 2061.7 2006.5 2054.7 2007.4 2048.9 C -f -S -n -1927.2 2062.4 m -1925.8 2060.1 1928.1 2058.2 1927 2056.4 C -1927.3 2055.5 1926.5 2053.5 1926.8 2051.8 C -1926.8 2052.8 1926 2052.5 1925.3 2052.5 C -1924.1 2052.8 1925 2050.5 1924.4 2050.1 C -1925.3 2050.2 1925.4 2048.8 1926.3 2049.4 C -1926.5 2052.3 1928.4 2047.2 1928.4 2051.1 C -1928.9 2050.5 1929 2051.4 1928.9 2051.8 C -1928.9 2052 1928.9 2052.3 1928.9 2052.5 C -1929.4 2051.4 1928.9 2049 1930.1 2048.2 C -1928.9 2047.1 1930.5 2047.1 1930.4 2046.5 C -1931.9 2046.2 1933.1 2046.1 1934.7 2046.5 C -1934.6 2046.9 1935.2 2047.9 1934.4 2048.4 C -1936.9 2048.1 1933.6 2043.8 1935.9 2043.9 C -1935.7 2043.9 1934.8 2041.3 1933.2 2041.7 C -1932.5 2041.6 1932.4 2039.6 1932.3 2041 C -1930.8 2042.6 1929 2040.6 1927.7 2042 C -1927.5 2041.4 1927.1 2040.9 1927.2 2040.3 C -1927.8 2040.6 1927.4 2039.1 1928.2 2038.6 C -1929.4 2038 1930.5 2038.8 1931.3 2037.9 C -1931.7 2039 1932.5 2038.6 1931.8 2037.6 C -1930.9 2037 1928.7 2037.8 1928.2 2037.9 C -1926.7 2037.8 1928 2039 1927 2038.8 C -1927.4 2040.4 1925.6 2040.8 1925.1 2041 C -1924.3 2040.4 1923.2 2040.5 1922.2 2040.5 C -1921.4 2041.7 1921 2043.9 1919.3 2043.9 C -1918.8 2043.4 1917.2 2043.3 1916.4 2043.4 C -1915.9 2044.4 1915.7 2046 1914.3 2046.5 C -1913.1 2046.6 1912 2044.5 1911.4 2046.3 C -1912.8 2046.5 1913.8 2047.4 1915.7 2047 C -1916.9 2047.7 1915.6 2048.8 1916 2049.4 C -1915.4 2049.3 1913.9 2050.3 1913.3 2051.1 C -1913.9 2054.1 1916 2050.2 1916.7 2053 C -1916.9 2053.8 1915.5 2054.1 1916.7 2054.4 C -1917 2054.7 1920.2 2054.3 1919.3 2056.6 C -1918.8 2056.1 1920.2 2058.6 1920.3 2057.6 C -1921.2 2057.9 1922.1 2057.5 1922.4 2059 C -1922.3 2059.1 1922.2 2059.3 1922 2059.2 C -1922.1 2059.7 1922.4 2060.3 1922.9 2060.7 C -1923.2 2060.1 1923.8 2060.4 1924.6 2060.7 C -1925.9 2062.6 1923.2 2062 1925.6 2063.6 C -1926.1 2063.1 1927.3 2062.5 1927.2 2062.4 C -[0.21 0.21 0 0] vc -f -S -n -1933.2 2063.3 m -1933.2 2060.7 1933.2 2058.2 1933.2 2055.6 C -1933.2 2058.2 1933.2 2060.7 1933.2 2063.3 C -[0 1 1 0.36] vc -f -S -n -1965.2 2049.2 m -1967.1 2050.1 1969.9 2053.7 1972.1 2056.4 C -1970.5 2054 1967.6 2051.3 1965.2 2049.2 C -f -S -n -1991.8 2034.8 m -1991.7 2041.5 1992 2048.5 1991.6 2055.2 C -1990.5 2056.4 1991.9 2054.9 1991.8 2054.4 C -1991.8 2047.9 1991.8 2041.3 1991.8 2034.8 C -f -S -n -1988.9 2053.2 m -1988.9 2044.3 1988.9 2036.6 1988.9 2028.3 C -1985.7 2028.2 1987.2 2023.5 1983.9 2024.2 C -1983.9 2022.4 1982 2021.6 1981 2021.3 C -1980.6 2021.1 1980.6 2021.7 1980.3 2021.6 C -1980.3 2027 1980.3 2034.8 1980.3 2041.5 C -1979.3 2043.2 1977.6 2043 1976.2 2043.6 C -1977.1 2043.8 1978.5 2043.2 1978.8 2044.1 C -1978.5 2045.3 1979.9 2045.3 1980.3 2045.8 C -1980.5 2046.8 1980.7 2046.2 1981.5 2046.5 C -1982.4 2047.1 1982 2048.6 1982.7 2049.4 C -1984.2 2049.6 1984.6 2052.2 1986.8 2051.6 C -1987.1 2048.6 1985.1 2042.7 1986.5 2040.5 C -1986.3 2036.7 1986.9 2031.7 1986 2029.2 C -1986.3 2027.1 1986.9 2028.6 1987.7 2027.6 C -1987.7 2028.3 1988.7 2028 1988.7 2028.8 C -1988.1 2033 1988.7 2037.5 1988.2 2041.7 C -1987.8 2041.4 1988 2040.8 1988 2040.3 C -1988 2041 1988 2042.4 1988 2042.4 C -1988 2042.4 1988.1 2042.3 1988.2 2042.2 C -1989.3 2046 1988 2050.2 1988.4 2054 C -1987.8 2054.4 1987.1 2054.7 1986.5 2055.4 C -1987.4 2054.4 1988.4 2054.6 1988.9 2053.2 C -[0 1 1 0.23] vc -f -S -n -1950.8 2054.4 m -1949.7 2053.4 1948.7 2052.3 1947.6 2051.3 C -1948.7 2052.3 1949.7 2053.4 1950.8 2054.4 C -[0 1 1 0.36] vc -f -S -n -vmrs -2006.7 2043.2 m -2004.5 2040.8 2002.4 2038.4 2000.2 2036 C -2002.4 2038.4 2004.5 2040.8 2006.7 2043.2 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -S -n -1976.7 2019.6 m -1975.8 2018.6 1976.4 2016.9 1976.2 2015.6 C -1976 2021.3 1975.8 2031.2 1976.2 2038.8 C -1976.4 2032.4 1975.8 2025.5 1976.7 2019.6 C -f -S -n -1988.4 2053.5 m -1988.6 2049.2 1988.1 2042.8 1988 2040 C -1988.4 2040.4 1988.1 2041 1988.2 2041.5 C -1988.3 2037.2 1988 2032.7 1988.4 2028.5 C -1987.6 2027.1 1987.2 2028.6 1986.8 2028 C -1985.9 2028.5 1986.5 2029.7 1986.3 2030.4 C -1986.9 2029.8 1986.6 2031 1987 2031.2 C -1987.4 2039.6 1985 2043 1987.2 2050.4 C -1987.2 2051.6 1985.9 2052.3 1984.6 2051.3 C -1981.9 2049.7 1982.9 2047 1980.3 2046.5 C -1980.3 2045.2 1978.1 2046.2 1978.6 2043.9 C -1975.6 2043.3 1979.3 2045.6 1979.6 2046.5 C -1980.8 2046.6 1981.5 2048.5 1982.2 2049.9 C -1983.7 2050.8 1984.8 2052.8 1986.5 2053 C -1986.7 2053.5 1987.5 2054.1 1987 2054.7 C -1987.4 2053.9 1988.3 2054.3 1988.4 2053.5 C -[0 1 1 0.23] vc -f -S -n -1988 2038.1 m -1988 2036.7 1988 2035.4 1988 2034 C -1988 2035.4 1988 2036.7 1988 2038.1 C -[0 1 1 0.36] vc -f -S -n -1999.7 2035.7 m -1997.6 2033.5 1995.4 2031.2 1993.2 2029 C -1995.4 2031.2 1997.6 2033.5 1999.7 2035.7 C -f -S -n -1944 2029.2 m -1945.2 2029.5 1946.9 2029.8 1947.9 2029 C -1947.4 2027.4 1949 2026.7 1949.8 2026.4 C -1953.9 2028.6 1952.6 2023.4 1955.6 2023.7 C -1957.4 2021.4 1960.7 2027 1959.4 2021.3 C -1959.3 2021.7 1959.6 2022.3 1959.2 2022.5 C -1958.1 2021.2 1960.1 2021.1 1959.4 2019.6 C -1959.1 2012.7 1959.9 2005.1 1959.6 1999.2 C -1955.3 2000.1 1951.3 2003.1 1947.2 2005 C -1943.9 2006 1941.2 2008.7 1938 2010 C -1936.9 2012.1 1938.2 2014.8 1937.8 2017.5 C -1937.8 2022.6 1937.8 2027.3 1937.8 2032.1 C -1938.2 2026.5 1938 2019 1938 2012.4 C -1938.5 2012 1939.2 2012.3 1939.7 2012.2 C -1940.8 2012.8 1939.7 2013.6 1939.7 2014.4 C -1940.4 2020.5 1939.4 2028 1939.7 2032.1 C -1940.6 2031.9 1941.2 2032.7 1941.9 2032.1 C -1941.7 2031.2 1943 2029.7 1944 2029.2 C -[0 0.2 1 0] vc -f -S -n -1937.1 2031.6 m -1937.1 2029.1 1937.1 2026.5 1937.1 2024 C -1937.1 2026.5 1937.1 2029.1 1937.1 2031.6 C -[0 1 1 0.36] vc -f -S -n -1991.8 2028 m -1992.5 2027.8 1993.2 2029.9 1994 2030.2 C -1992.9 2029.6 1993.1 2028.1 1991.8 2028 C -[0 1 1 0.23] vc -f -S -n -1991.8 2027.8 m -1992.4 2027.6 1992.6 2028.3 1993 2028.5 C -1992.6 2028.2 1992.2 2027.6 1991.6 2027.8 C -1991.6 2028.5 1991.6 2029.1 1991.6 2029.7 C -1991.6 2029.1 1991.4 2028.3 1991.8 2027.8 C -[0 1 1 0.36] vc -f -S -n -1985.8 2025.4 m -1985.3 2025.2 1984.8 2024.7 1984.1 2024.9 C -1983.3 2025.3 1983.6 2027.3 1983.9 2027.6 C -1985 2028 1986.9 2026.9 1985.8 2025.4 C -[0 1 1 0.23] vc -f -S -n -vmrs -1993.5 2024.4 m -1992.4 2023.7 1991.3 2022.9 1990.1 2023.2 C -1990.7 2023.7 1989.8 2023.8 1989.4 2023.7 C -1989.1 2023.7 1988.6 2023.9 1988.4 2023.5 C -1988.5 2023.2 1988.3 2022.7 1988.7 2022.5 C -1989 2022.6 1988.9 2023 1988.9 2023.2 C -1989.1 2022.8 1990.4 2022.3 1990.6 2021.3 C -1990.4 2021.8 1990 2021.3 1990.1 2021.1 C -1990.1 2020.9 1990.1 2020.1 1990.1 2020.6 C -1989.9 2021.1 1989.5 2020.6 1989.6 2020.4 C -1989.6 2019.8 1988.7 2019.6 1988.2 2019.2 C -1987.5 2018.7 1987.7 2020.2 1987 2019.4 C -1987.5 2020.4 1986 2021.1 1987.5 2021.8 C -1986.8 2023.1 1986.6 2021.1 1986 2021.1 C -1986.1 2020.1 1985.9 2019 1986.3 2018.2 C -1986.7 2018.4 1986.5 2019 1986.5 2019.4 C -1986.5 2018.7 1986.4 2017.8 1987.2 2017.7 C -1986.5 2017.2 1985.5 2019.3 1985.3 2020.4 C -1986.2 2022 1987.3 2023.5 1989.2 2024.2 C -1990.8 2024.3 1991.6 2022.9 1993.2 2024.4 C -1993.8 2025.4 1995 2026.6 1995.9 2027.1 C -1995 2026.5 1994.1 2025.5 1993.5 2024.4 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -[0 0.5 0.5 0.2] vc -S -n -2023 2040.3 m -2023.2 2036 2022.7 2029.6 2022.5 2026.8 C -2022.9 2027.2 2022.7 2027.8 2022.8 2028.3 C -2022.8 2024 2022.6 2019.5 2023 2015.3 C -2022.2 2013.9 2021.7 2015.4 2021.3 2014.8 C -2020.4 2015.3 2021 2016.5 2020.8 2017.2 C -2021.4 2016.6 2021.1 2017.8 2021.6 2018 C -2022 2026.4 2019.6 2029.8 2021.8 2037.2 C -2021.7 2038.4 2020.5 2039.1 2019.2 2038.1 C -2016.5 2036.5 2017.5 2033.8 2014.8 2033.3 C -2014.9 2032 2012.6 2033 2013.2 2030.7 C -2011.9 2030.8 2011.2 2030.1 2010.8 2029.2 C -2010.8 2029.1 2010.8 2028.2 2010.8 2028.8 C -2010 2028.8 2010.4 2026.5 2008.6 2027.3 C -2007.9 2026.6 2007.3 2025.9 2007.9 2027.1 C -2009.7 2028 2010 2030.1 2012.2 2030.9 C -2012.9 2032.1 2013.7 2033.6 2015.1 2033.6 C -2015.7 2035.1 2016.9 2036.7 2018.4 2038.4 C -2019.8 2039.3 2022 2039.4 2021.6 2041.5 C -2021.9 2040.7 2022.9 2041.1 2023 2040.3 C -[0 1 1 0.23] vc -f -S -n -2022.5 2024.9 m -2022.5 2023.5 2022.5 2022.2 2022.5 2020.8 C -2022.5 2022.2 2022.5 2023.5 2022.5 2024.9 C -[0 1 1 0.36] vc -f -S -n -1983.2 2022.8 m -1982.4 2022.5 1982.1 2021.6 1981.2 2022.3 C -1981.1 2022.9 1980.5 2024 1981 2024.2 C -1981.8 2024.6 1982.9 2024.4 1983.2 2022.8 C -[0 1 1 0.23] vc -f -S -n -1931.1 2019.9 m -1929.6 2017.7 1932 2015.7 1930.8 2013.9 C -1931.1 2013 1930.3 2011 1930.6 2009.3 C -1930.6 2010.3 1929.8 2010 1929.2 2010 C -1928 2010.3 1928.8 2008.1 1928.2 2007.6 C -1929.1 2007.8 1929.3 2006.3 1930.1 2006.9 C -1930.3 2009.8 1932.2 2004.8 1932.3 2008.6 C -1932.7 2008 1932.8 2009 1932.8 2009.3 C -1932.8 2009.6 1932.8 2009.8 1932.8 2010 C -1933.2 2009 1932.7 2006.6 1934 2005.7 C -1932.7 2004.6 1934.3 2004.6 1934.2 2004 C -1935.8 2003.7 1937 2003.6 1938.5 2004 C -1938.5 2004.5 1939.1 2005.4 1938.3 2006 C -1940.7 2005.7 1937.4 2001.3 1939.7 2001.4 C -1939.5 2001.4 1938.6 1998.8 1937.1 1999.2 C -1936.3 1999.1 1936.2 1997.1 1936.1 1998.5 C -1934.7 2000.1 1932.9 1998.2 1931.6 1999.5 C -1931.3 1998.9 1930.9 1998.5 1931.1 1997.8 C -1931.6 1998.2 1931.3 1996.6 1932 1996.1 C -1933.2 1995.5 1934.3 1996.4 1935.2 1995.4 C -1935.5 1996.5 1936.3 1996.1 1935.6 1995.2 C -1934.7 1994.5 1932.5 1995.3 1932 1995.4 C -1930.5 1995.3 1931.9 1996.5 1930.8 1996.4 C -1931.2 1997.9 1929.5 1998.3 1928.9 1998.5 C -1928.1 1997.9 1927.1 1998 1926 1998 C -1925.3 1999.2 1924.8 2001.4 1923.2 2001.4 C -1922.6 2000.9 1921 2000.9 1920.3 2000.9 C -1919.7 2001.9 1919.6 2003.5 1918.1 2004 C -1916.9 2004.1 1915.8 2002 1915.2 2003.8 C -1916.7 2004 1917.6 2004.9 1919.6 2004.5 C -1920.7 2005.2 1919.4 2006.3 1919.8 2006.9 C -1919.2 2006.9 1917.7 2007.8 1917.2 2008.6 C -1917.8 2011.6 1919.8 2007.8 1920.5 2010.5 C -1920.8 2011.3 1919.3 2011.6 1920.5 2012 C -1920.8 2012.3 1924 2011.8 1923.2 2014.1 C -1922.6 2013.6 1924.1 2016.1 1924.1 2015.1 C -1925.1 2015.4 1925.9 2015 1926.3 2016.5 C -1926.2 2016.6 1926 2016.8 1925.8 2016.8 C -1925.9 2017.2 1926.2 2017.8 1926.8 2018.2 C -1927.1 2017.6 1927.7 2018 1928.4 2018.2 C -1929.7 2020.1 1927.1 2019.5 1929.4 2021.1 C -1929.9 2020.7 1931.1 2020 1931.1 2019.9 C -[0.21 0.21 0 0] vc -f -S -n -1937.1 2020.8 m -1937.1 2018.3 1937.1 2015.7 1937.1 2013.2 C -1937.1 2015.7 1937.1 2018.3 1937.1 2020.8 C -[0 1 1 0.36] vc -f -S -n -2020.4 2012.2 m -2019.8 2012 2019.3 2011.5 2018.7 2011.7 C -2017.9 2012.1 2018.1 2014.1 2018.4 2014.4 C -2019.6 2014.8 2021.4 2013.7 2020.4 2012.2 C -[0 1 1 0.23] vc -f -S -n -1976 2013.9 m -1973.8 2011.5 1971.6 2009.1 1969.5 2006.7 C -1971.6 2009.1 1973.8 2011.5 1976 2013.9 C -[0 1 1 0.36] vc -f -S -n -1995.4 2012.7 m -1996.1 2010.3 1993.8 2006.2 1997.3 2005.7 C -1998.9 2005.4 2000 2003.7 2001.4 2003.1 C -2003.9 2003.1 2005.3 2001.3 2006.9 1999.7 C -2004.5 2003.5 2000 2002.2 1997.6 2005.7 C -1996.5 2005.9 1994.8 2006.1 1995.2 2007.6 C -1995.7 2009.4 1995.2 2011.6 1994.7 2012.9 C -1992 2015.8 1987.8 2015.7 1985.3 2018.7 C -1988.3 2016.3 1992.3 2015.3 1995.4 2012.7 C -[0.18 0.18 0 0.78] vc -f -S -n -1995.6 2012.4 m -1995.6 2011.2 1995.6 2010 1995.6 2008.8 C -1995.6 2010 1995.6 2011.2 1995.6 2012.4 C -[0 1 1 0.36] vc -f -S -n -vmrs -2017.7 2009.6 m -2016.9 2009.3 2016.7 2008.4 2015.8 2009.1 C -2014.2 2010.6 2016 2010.6 2016.5 2011.5 C -2017.2 2010.9 2018.1 2010.8 2017.7 2009.6 C -[0 1 1 0.23] vc -f -0.4 w -2 J -2 M -S -n -2014.4 2006.4 m -2013.5 2006.8 2012.1 2005.6 2012 2006.7 C -2013 2007.3 2011.9 2009.2 2012.9 2008.4 C -2014.2 2008.3 2014.6 2007.8 2014.4 2006.4 C -f -S -n -1969 2006.4 m -1966.5 2003.8 1964 2001.2 1961.6 1998.5 C -1964 2001.2 1966.5 2003.8 1969 2006.4 C -[0 1 1 0.36] vc -f -S -n -2012 2005.2 m -2012.2 2004.2 2011.4 2003.3 2010.3 2003.3 C -2009 2003.6 2010 2004.7 2009.6 2004.8 C -2009.3 2005.7 2011.4 2006.7 2012 2005.2 C -[0 1 1 0.23] vc -f -S -n -1962.8 1995.2 m -1961.7 1994.4 1960.6 1993.7 1959.4 1994 C -1959.5 1994.9 1957.5 1994.1 1956.8 1994.7 C -1955.9 1995.5 1956.7 1997 1955.1 1997.3 C -1956.9 1996.7 1956.8 1994 1959.2 1994.7 C -1961.1 1991 1968.9 2003.2 1962.8 1995.2 C -[0 1 1 0.36] vc -f -S -n -1954.6 1995.6 m -1955.9 1994.7 1955.1 1989.8 1955.3 1988 C -1954.5 1988.3 1954.9 1986.6 1954.4 1986 C -1955.7 1989.2 1953.9 1991.1 1954.8 1994.2 C -1954.5 1995.9 1953.5 1995.3 1953.9 1997.3 C -1955.3 1998.3 1953.2 1995.5 1954.6 1995.6 C -f -S -n -1992.3 2011 m -1992.5 2006.7 1992 2000.3 1991.8 1997.6 C -1992.2 1997.9 1992 1998.5 1992 1999 C -1992.1 1994.7 1991.9 1990.2 1992.3 1986 C -1991.4 1984.6 1991 1986.1 1990.6 1985.6 C -1989.7 1986 1990.3 1987.2 1990.1 1988 C -1990.7 1987.4 1990.4 1988.5 1990.8 1988.7 C -1991.3 1997.1 1988.9 2000.6 1991.1 2007.9 C -1991 2009.1 1989.8 2009.9 1988.4 2008.8 C -1985.7 2007.2 1986.8 2004.5 1984.1 2004 C -1984.2 2002.7 1981.9 2003.7 1982.4 2001.4 C -1981.2 2001.5 1980.5 2000.8 1980 2000 C -1980 1999.8 1980 1998.9 1980 1999.5 C -1979.3 1999.5 1979.7 1997.2 1977.9 1998 C -1977.2 1997.3 1976.6 1996.7 1977.2 1997.8 C -1979 1998.7 1979.3 2000.8 1981.5 2001.6 C -1982.2 2002.8 1983 2004.3 1984.4 2004.3 C -1985 2005.8 1986.2 2007.5 1987.7 2009.1 C -1989 2010 1991.3 2010.2 1990.8 2012.2 C -1991.2 2011.4 1992.2 2011.8 1992.3 2011 C -[0 1 1 0.23] vc -f -S -n -1991.8 1995.6 m -1991.8 1994.3 1991.8 1992.9 1991.8 1991.6 C -1991.8 1992.9 1991.8 1994.3 1991.8 1995.6 C -[0 1 1 0.36] vc -f -S -n -1959.2 1994.2 m -1958.8 1993.3 1960.7 1993.9 1961.1 1993.7 C -1961.5 1993.9 1961.2 1994.4 1961.8 1994.2 C -1960.9 1994 1960.8 1992.9 1959.9 1992.5 C -1959.6 1993.5 1958.3 1993.5 1958.2 1994.2 C -1958.1 1994.1 1958 1994 1958 1994 C -1957.2 1994.9 1958 1993.4 1956.8 1993 C -1955.6 1992.5 1956 1991 1956.3 1989.9 C -1956.5 1989.8 1956.6 1990 1956.8 1990.1 C -1957.1 1989 1956 1989.1 1955.8 1988.2 C -1955.1 1990.4 1956.2 1995 1954.8 1995.9 C -1954.1 1995.5 1954.5 1996.5 1954.4 1997.1 C -1955 1996.8 1954.8 1997.4 1955.6 1996.8 C -1956 1996 1956.3 1993.2 1958.7 1994.2 C -1958.9 1994.2 1959.7 1994.2 1959.2 1994.2 C -[0 1 1 0.23] vc -f -S -n -1958.2 1994 m -1958.4 1993.5 1959.7 1993.1 1959.9 1992 C -1959.7 1992.5 1959.3 1992 1959.4 1991.8 C -1959.4 1991.6 1959.4 1990.8 1959.4 1991.3 C -1959.2 1991.8 1958.8 1991.3 1958.9 1991.1 C -1958.9 1990.5 1958 1990.3 1957.5 1989.9 C -1956.8 1989.5 1956.9 1991 1956.3 1990.1 C -1956.7 1991 1955.4 1992.1 1956.5 1992.3 C -1956.8 1993.5 1958.3 1992.9 1957.2 1994 C -1957.8 1994.3 1958.1 1992.4 1958.2 1994 C -[0 0.5 0.5 0.2] vc -f -S -n -vmrs -1954.4 1982.7 m -1956.1 1982.7 1954.1 1982.5 1953.9 1982.9 C -1953.9 1983.7 1953.7 1984.7 1954.1 1985.3 C -1954.4 1984.2 1953.6 1983.6 1954.4 1982.7 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -S -n -1989.6 1982.9 m -1989.1 1982.7 1988.6 1982.3 1988 1982.4 C -1987.2 1982.8 1987.4 1984.8 1987.7 1985.1 C -1988.9 1985.6 1990.7 1984.4 1989.6 1982.9 C -[0 1 1 0.23] vc -f -S -n -1987 1980.3 m -1986.2 1980 1986 1979.1 1985.1 1979.8 C -1983.5 1981.4 1985.3 1981.4 1985.8 1982.2 C -1986.5 1981.7 1987.4 1981.5 1987 1980.3 C -f -S -n -1983.6 1977.2 m -1982.7 1977.5 1981.4 1976.3 1981.2 1977.4 C -1982.3 1978 1981.2 1979.9 1982.2 1979.1 C -1983.5 1979 1983.9 1978.5 1983.6 1977.2 C -f -S -n -1981.2 1976 m -1981.5 1974.9 1980.6 1974 1979.6 1974 C -1978.3 1974.3 1979.3 1975.4 1978.8 1975.5 C -1978.6 1976.4 1980.7 1977.4 1981.2 1976 C -f -S -n -1972.1 2082.3 m -1971.8 2081.8 1971.3 2080.9 1971.2 2080.1 C -1971.1 2072.9 1971.3 2064.6 1970.9 2058.3 C -1970.3 2058.5 1970.1 2057.7 1969.7 2058.5 C -1970.6 2058.5 1969.7 2059 1970.2 2059.2 C -1970.2 2065.4 1970.2 2072.4 1970.2 2077.7 C -1971.1 2078.9 1970.6 2078.9 1970.4 2079.9 C -1969.2 2080.2 1968.2 2080.4 1967.3 2079.6 C -1966.8 2077.8 1963.4 2076.3 1963.5 2075.1 C -1961.5 2075.5 1962 2071.5 1959.6 2072 C -1959.2 2070 1956.5 2069.3 1955.8 2067.6 C -1956 2068.4 1955.3 2069.7 1956.5 2069.8 C -1958.6 2068.9 1958.1 2073.5 1960.1 2072.4 C -1960.7 2075.9 1964.7 2074.6 1964.2 2078 C -1967.2 2078.6 1967.9 2081.6 1970.7 2080.6 C -1970.3 2081.1 1971.5 2081.2 1971.9 2082.3 C -1967.2 2084.3 1962.9 2087.1 1958.2 2089 C -1962.9 2087 1967.4 2084.4 1972.1 2082.3 C -[0 0.2 1 0] vc -f -S -n -1971.9 2080.1 m -1971.9 2075.1 1971.9 2070 1971.9 2065 C -1971.9 2070 1971.9 2075.1 1971.9 2080.1 C -[0 1 1 0.23] vc -f -S -n -2010.8 2050.6 m -2013.2 2049 2010.5 2050.1 2010.5 2051.3 C -2010.5 2057.7 2010.5 2064.1 2010.5 2070.5 C -2008.7 2072.4 2006 2073.3 2003.6 2074.4 C -2016.4 2073.7 2008 2058.4 2010.8 2050.6 C -[0.4 0.4 0 0] vc -f -S -n -2006.4 2066.9 m -2006.4 2061.9 2006.4 2056.8 2006.4 2051.8 C -2006.4 2056.8 2006.4 2061.9 2006.4 2066.9 C -[0 1 1 0.23] vc -f -S -n -1971.9 2060.7 m -1972.2 2060.3 1971.4 2068.2 1972.4 2061.9 C -1971.8 2061.6 1972.4 2060.9 1971.9 2060.7 C -f -S -n -vmrs -1986.5 2055.2 m -1987.5 2054.3 1986.3 2053.4 1986 2052.8 C -1983.8 2052.7 1983.6 2050.1 1981.7 2049.6 C -1981.2 2048.7 1980.8 2047 1980.3 2046.8 C -1978.5 2047 1978 2044.6 1976.7 2043.9 C -1974 2044.4 1972 2046.6 1969.2 2047 C -1969 2047.2 1968.8 2047.5 1968.5 2047.7 C -1970.6 2049.6 1973.1 2051.3 1974.3 2054.2 C -1975.7 2054.5 1977 2055.2 1976.4 2057.1 C -1976.7 2058 1975.5 2058.5 1976 2059.5 C -1979.2 2058 1983 2056.6 1986.5 2055.2 C -[0 0.5 0.5 0.2] vc -f -0.4 w -2 J -2 M -S -n -1970.2 2054.2 m -1971.5 2055.3 1972.5 2056.8 1972.1 2058.3 C -1972.8 2056.5 1971.6 2055.6 1970.2 2054.2 C -[0 1 1 0.23] vc -f -S -n -1992 2052.5 m -1992 2053.4 1992.2 2054.4 1991.8 2055.2 C -1992.2 2054.4 1992 2053.4 1992 2052.5 C -f -S -n -1957.2 2053 m -1958.1 2052.6 1959 2052.2 1959.9 2051.8 C -1959 2052.2 1958.1 2052.6 1957.2 2053 C -f -S -n -2006.4 2047.5 m -2006.8 2047.1 2006 2055 2006.9 2048.7 C -2006.4 2048.4 2007 2047.7 2006.4 2047.5 C -f -S -n -2004.8 2041 m -2006.1 2042.1 2007.1 2043.6 2006.7 2045.1 C -2007.3 2043.3 2006.2 2042.4 2004.8 2041 C -f -S -n -1976 2039.8 m -1975.6 2039.3 1975.2 2038.4 1975 2037.6 C -1974.9 2030.4 1975.2 2022.1 1974.8 2015.8 C -1974.2 2016 1974 2015.3 1973.6 2016 C -1974.4 2016 1973.5 2016.5 1974 2016.8 C -1974 2022.9 1974 2030 1974 2035.2 C -1974.9 2036.4 1974.4 2036.4 1974.3 2037.4 C -1973.1 2037.7 1972 2037.9 1971.2 2037.2 C -1970.6 2035.3 1967.3 2033.9 1967.3 2032.6 C -1965.3 2033 1965.9 2029.1 1963.5 2029.5 C -1963 2027.6 1960.4 2026.8 1959.6 2025.2 C -1959.8 2025.9 1959.2 2027.2 1960.4 2027.3 C -1962.5 2026.4 1961.9 2031 1964 2030 C -1964.6 2033.4 1968.5 2032.1 1968 2035.5 C -1971 2036.1 1971.8 2039.1 1974.5 2038.1 C -1974.2 2038.7 1975.3 2038.7 1975.7 2039.8 C -1971 2041.8 1966.7 2044.6 1962 2046.5 C -1966.8 2044.5 1971.3 2041.9 1976 2039.8 C -[0 0.2 1 0] vc -f -S -n -1975.7 2037.6 m -1975.7 2032.6 1975.7 2027.6 1975.7 2022.5 C -1975.7 2027.6 1975.7 2032.6 1975.7 2037.6 C -[0 1 1 0.23] vc -f -S -n -1992 2035.5 m -1992 2034.2 1992 2032.9 1992 2031.6 C -1992 2032.9 1992 2034.2 1992 2035.5 C -f -S -n -2015.3 2036 m -2015.4 2034.1 2013.3 2034 2012.9 2033.3 C -2011.5 2031 2009.3 2029.4 2007.4 2028 C -2006.9 2027.1 2006.6 2023.8 2005 2024.9 C -2004 2024.9 2002.9 2024.9 2001.9 2024.9 C -2001.4 2026.5 2001 2028.4 2003.8 2028.3 C -2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C -2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C -2013 2035.5 2015.3 2037.4 2015.3 2036 C -[0 0 0 0] vc -f -S -n -vmrs -2009.1 2030.4 m -2009.1 2029 2007.5 2029.4 2006.9 2028.3 C -2007.2 2027.1 2006.5 2025.5 2005.7 2024.7 C -2004.6 2025.1 2003.1 2024.9 2001.9 2024.9 C -2001.8 2026.2 2000.9 2027 2002.4 2028 C -2004.5 2027.3 2004.9 2029.4 2006.9 2029 C -2007 2030.2 2007.6 2030.7 2008.4 2031.4 C -2008.8 2031.5 2009.1 2031.1 2009.1 2030.4 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -2003.8 2029.5 m -2003 2029.4 2001.9 2029.1 2002.4 2030.4 C -2003.1 2031.3 2005.2 2030.3 2003.8 2029.5 C -[0 1 1 0.23] vc -f -S -n -1999.2 2025.2 m -1999.1 2025.6 1998 2025.7 1998.8 2026.6 C -2000.9 2028.5 1999.5 2023.4 1999.2 2025.2 C -f -S -n -2007.6 2024.2 m -2007.6 2022.9 2008.4 2024.2 2007.6 2022.8 C -2007.6 2017.5 2007.8 2009.1 2007.4 2003.8 C -2007.9 2003.7 2008.7 2002.8 2009.1 2002.1 C -2009.6 2000.8 2008.3 2000.8 2007.9 2000.2 C -2004.9 2000 2008.9 2001.3 2007.2 2002.1 C -2006.7 2007.7 2007 2015.1 2006.9 2021.1 C -2006.7 2022.1 2005.4 2022.8 2006.2 2023.5 C -2006.6 2023.1 2008 2025.9 2007.6 2024.2 C -f -S -n -1989.9 2023.5 m -1989.5 2022.6 1991.4 2023.2 1991.8 2023 C -1992.2 2023.2 1991.9 2023.7 1992.5 2023.5 C -1991.6 2023.2 1991.6 2022.2 1990.6 2021.8 C -1990.4 2022.8 1989 2022.8 1988.9 2023.5 C -1988.5 2023 1988.7 2022.6 1988.7 2023.5 C -1989.1 2023.5 1990.2 2023.5 1989.9 2023.5 C -f -[0 0.5 0.5 0.2] vc -S -n -2003.3 2023.5 m -2003.1 2023.3 2003.1 2023.2 2003.3 2023 C -2003.7 2023.1 2003.9 2022.9 2003.8 2022.5 C -2003.4 2022.2 2001.2 2022.3 2002.4 2023 C -2002.6 2022.9 2002.7 2023.1 2002.8 2023.2 C -2000.7 2023.7 2003.9 2023.4 2003.3 2023.5 C -[0 1 1 0.23] vc -f -S -n -1986.8 2019.4 m -1987.8 2019.8 1987.5 2018.6 1987.2 2018 C -1986.2 2017.8 1987.3 2020.5 1986.3 2019.2 C -1986.3 2017.7 1986.3 2020.6 1986.3 2021.3 C -1988.5 2023.1 1985.6 2020.3 1986.8 2019.4 C -f -S -n -1975.7 2018.2 m -1976.1 2017.8 1975.2 2025.7 1976.2 2019.4 C -1975.7 2019.2 1976.3 2018.4 1975.7 2018.2 C -f -S -n -1974 2011.7 m -1975.4 2012.8 1976.4 2014.3 1976 2015.8 C -1976.6 2014 1975.5 2013.1 1974 2011.7 C -f -S -n -1984.6 2006.7 m -1984.7 2004.8 1982.6 2004.8 1982.2 2004 C -1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C -1976.1 1997.8 1975.8 1994.5 1974.3 1995.6 C -1973.3 1995.6 1972.2 1995.6 1971.2 1995.6 C -1970.7 1997.2 1970.3 1999.1 1973.1 1999 C -1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C -1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C -1982.3 2006.2 1984.5 2008.1 1984.6 2006.7 C -[0 0 0 0] vc -f -S -n -vmrs -1978.4 2001.2 m -1978.4 1999.7 1976.8 2000.1 1976.2 1999 C -1976.5 1997.8 1975.8 1996.2 1975 1995.4 C -1973.9 1995.8 1972.4 1995.6 1971.2 1995.6 C -1971 1997 1970.2 1997.7 1971.6 1998.8 C -1973.8 1998 1974.2 2000.1 1976.2 1999.7 C -1976.3 2000.9 1976.9 2001.4 1977.6 2002.1 C -1978.1 2002.2 1978.4 2001.8 1978.4 2001.2 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1973.1 2000.2 m -1972.3 2000.1 1971.2 1999.8 1971.6 2001.2 C -1972.4 2002 1974.5 2001 1973.1 2000.2 C -[0 1 1 0.23] vc -f -S -n -1960.8 1998.5 m -1961.6 1998.2 1962.6 2000.3 1963.2 2000.9 C -1962.3 2000.1 1962.2 1998.7 1960.8 1998.5 C -f -S -n -1968.5 1995.9 m -1968.4 1996.4 1967.3 1996.4 1968 1997.3 C -1970.1 1999.2 1968.8 1994.1 1968.5 1995.9 C -f -S -n -1976.9 1994.9 m -1976.9 1993.7 1977.6 1994.9 1976.9 1993.5 C -1976.9 1988.2 1977.1 1979.8 1976.7 1974.5 C -1977.2 1974.5 1978 1973.5 1978.4 1972.8 C -1978.8 1971.5 1977.6 1971.5 1977.2 1970.9 C -1974.2 1970.7 1978.2 1972 1976.4 1972.8 C -1976 1978.4 1976.3 1985.8 1976.2 1991.8 C -1976 1992.8 1974.6 1993.5 1975.5 1994.2 C -1975.9 1993.8 1977.3 1996.6 1976.9 1994.9 C -f -S -n -1972.6 1994.2 m -1972.4 1994 1972.4 1993.9 1972.6 1993.7 C -1973 1993.8 1973.1 1993.7 1973.1 1993.2 C -1972.7 1992.9 1970.5 1993.1 1971.6 1993.7 C -1971.9 1993.7 1972 1993.8 1972.1 1994 C -1970 1994.4 1973.1 1994.1 1972.6 1994.2 C -f -S -n -1948.1 2093.8 m -1947 2092.7 1945.9 2091.6 1944.8 2090.4 C -1945.9 2091.6 1947 2092.7 1948.1 2093.8 C -[0 0.4 1 0] vc -f -S -n -1953.4 2091.4 m -1954.8 2090.7 1956.3 2090 1957.7 2089.2 C -1956.3 2090 1954.8 2090.7 1953.4 2091.4 C -[0 0.2 1 0] vc -f -S -n -1954.1 2091.4 m -1956.6 2089.6 1957.2 2089.6 1954.1 2091.4 C -[0 0.4 1 0] vc -f -S -n -1962.3 2087.3 m -1963.7 2086.6 1965.2 2085.9 1966.6 2085.2 C -1965.2 2085.9 1963.7 2086.6 1962.3 2087.3 C -f -S -n -vmrs -1967.1 2084.9 m -1968.3 2084.4 1969.7 2083.8 1970.9 2083.2 C -1969.7 2083.8 1968.3 2084.4 1967.1 2084.9 C -[0 0.4 1 0] vc -f -0.4 w -2 J -2 M -S -n -1982.7 2080.6 m -1981.5 2079.5 1980.5 2078.4 1979.3 2077.2 C -1980.5 2078.4 1981.5 2079.5 1982.7 2080.6 C -f -S -n -1988 2078.2 m -1989.4 2077.5 1990.8 2076.8 1992.3 2076 C -1990.8 2076.8 1989.4 2077.5 1988 2078.2 C -[0 0.2 1 0] vc -f -S -n -1988.7 2078.2 m -1991.1 2076.4 1991.8 2076.4 1988.7 2078.2 C -[0 0.4 1 0] vc -f -S -n -1976.2 2063.8 m -1978.6 2062.2 1976 2063.3 1976 2064.5 C -1976.1 2067.8 1975.5 2071.4 1976.4 2074.4 C -1975.7 2071.1 1975.9 2067.2 1976.2 2063.8 C -f -S -n -1996.8 2074.1 m -1998.3 2073.4 1999.7 2072.7 2001.2 2072 C -1999.7 2072.7 1998.3 2073.4 1996.8 2074.1 C -f -S -n -2001.6 2071.7 m -2002.9 2071.2 2004.2 2070.6 2005.5 2070 C -2004.2 2070.6 2002.9 2071.2 2001.6 2071.7 C -f -S -n -1981.5 2060.7 m -1980.2 2061.2 1978.9 2061.5 1977.9 2062.6 C -1978.9 2061.5 1980.2 2061.2 1981.5 2060.7 C -f -S -n -1982 2060.4 m -1982.7 2060.1 1983.6 2059.8 1984.4 2059.5 C -1983.6 2059.8 1982.7 2060.1 1982 2060.4 C -f -S -n -1952 2051.3 m -1950.8 2050.2 1949.7 2049.1 1948.6 2048 C -1949.7 2049.1 1950.8 2050.2 1952 2051.3 C -f -S -n -vmrs -1977.4 2047.7 m -1975.8 2047.8 1974.8 2046.1 1974.5 2045.3 C -1974.9 2044.4 1976 2044.5 1976.7 2044.8 C -1977.9 2045 1977 2048.4 1979.3 2047.5 C -1979.9 2047.5 1980.8 2048.6 1979.8 2049.2 C -1978.2 2050.4 1980.8 2049.5 1980.3 2049.4 C -1981.4 2049.8 1980.3 2048.4 1980.3 2048 C -1979.8 2047.5 1979 2046.6 1978.4 2046.5 C -1977.3 2045.9 1977.2 2043.3 1975.2 2044.6 C -1974.7 2045.3 1973.6 2045 1973.3 2045.8 C -1975 2046.3 1975.8 2049.8 1978.1 2049.4 C -1978.4 2050.9 1978.7 2048.5 1977.9 2049.2 C -1977.7 2048.7 1977.2 2047.8 1977.4 2047.7 C -[0 0.5 0.5 0.2] vc -f -0.4 w -2 J -2 M -S -n -1957.2 2048.9 m -1958.7 2048.2 1960.1 2047.5 1961.6 2046.8 C -1960.1 2047.5 1958.7 2048.2 1957.2 2048.9 C -[0 0.2 1 0] vc -f -S -n -1958 2048.9 m -1960.4 2047.1 1961.1 2047.1 1958 2048.9 C -[0 0.4 1 0] vc -f -S -n -1966.1 2044.8 m -1967.6 2044.1 1969 2043.4 1970.4 2042.7 C -1969 2043.4 1967.6 2044.1 1966.1 2044.8 C -f -S -n -1970.9 2042.4 m -1972.2 2041.9 1973.5 2041.3 1974.8 2040.8 C -1973.5 2041.3 1972.2 2041.9 1970.9 2042.4 C -f -S -n -2012 2034.5 m -2010.4 2034.6 2009.3 2032.9 2009.1 2032.1 C -2009.4 2031 2010.3 2031.3 2011.2 2031.6 C -2012.5 2031.8 2011.6 2035.2 2013.9 2034.3 C -2014.4 2034.3 2015.4 2035.4 2014.4 2036 C -2012.7 2037.2 2015.3 2036.3 2014.8 2036.2 C -2015.9 2036.6 2014.8 2035.2 2014.8 2034.8 C -2014.4 2034.3 2013.6 2033.4 2012.9 2033.3 C -2011.5 2031 2009.3 2029.4 2007.4 2028 C -2007.5 2026.5 2007.3 2027.9 2007.2 2028.3 C -2007.9 2028.8 2008.7 2029.1 2009.3 2030 C -2009.6 2030.7 2009 2031.9 2008.4 2031.6 C -2006.7 2031 2007.7 2028 2005 2028.8 C -2004.8 2028.6 2004.3 2028.2 2003.8 2028.3 C -2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C -2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C -2012.7 2036.1 2011.8 2035 2012 2034.5 C -[0 0.5 0.5 0.2] vc -f -S -n -1981.2 2005.2 m -1979.7 2005.3 1978.6 2003.6 1978.4 2002.8 C -1978.7 2001.8 1979.6 2002.1 1980.5 2002.4 C -1981.8 2002.5 1980.9 2005.9 1983.2 2005 C -1983.7 2005.1 1984.7 2006.1 1983.6 2006.7 C -1982 2007.9 1984.6 2007 1984.1 2006.9 C -1985.2 2007.3 1984.1 2006 1984.1 2005.5 C -1983.6 2005 1982.9 2004.1 1982.2 2004 C -1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C -1976.7 1997.2 1976.6 1998.6 1976.4 1999 C -1977.2 1999.5 1978 1999.8 1978.6 2000.7 C -1978.8 2001.5 1978.3 2002.7 1977.6 2002.4 C -1976 2001.8 1977 1998.7 1974.3 1999.5 C -1974.1 1999.3 1973.6 1998.9 1973.1 1999 C -1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C -1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C -1982 2006.8 1981.1 2005.7 1981.2 2005.2 C -f -S -n -1966.8 1976.4 m -1969.4 1973 1974.4 1974.6 1976.2 1970.4 C -1972.7 1974 1968 1975.1 1964 1977.4 C -1960.9 1979.9 1957.1 1981.8 1953.9 1982.7 C -1958.4 1981.1 1962.6 1978.8 1966.8 1976.4 C -[0.18 0.18 0 0.78] vc -f -S -n -1948.4 2093.8 m -1949.8 2093.1 1951.2 2092.5 1952.7 2091.9 C -1951.2 2092.5 1949.8 2093.1 1948.4 2093.8 C -[0 0.2 1 0] vc -f -S -n -1948.1 2093.6 m -1947.3 2092.8 1946.5 2091.9 1945.7 2091.2 C -1946.5 2091.9 1947.3 2092.8 1948.1 2093.6 C -f -S -n -vmrs -1942.1 2087.8 m -1943.5 2088.4 1944.3 2089.5 1945.2 2090.7 C -1944.8 2089.3 1943.3 2088.3 1942.1 2087.8 C -[0 0.2 1 0] vc -f -0.4 w -2 J -2 M -S -n -1933.5 2078.4 m -1933.5 2078 1933.2 2079 1933.7 2079.4 C -1935 2080.4 1936.2 2081.3 1937.1 2082.8 C -1936.7 2080.7 1933.7 2080.7 1933.5 2078.4 C -f -S -n -1982.9 2080.6 m -1984.4 2079.9 1985.8 2079.3 1987.2 2078.7 C -1985.8 2079.3 1984.4 2079.9 1982.9 2080.6 C -f -S -n -1982.7 2080.4 m -1981.9 2079.6 1981.1 2078.7 1980.3 2078 C -1981.1 2078.7 1981.9 2079.6 1982.7 2080.4 C -f -S -n -1977.4 2075.1 m -1977.9 2075.3 1979.1 2076.4 1979.8 2077.5 C -1979 2076.8 1978.7 2075.1 1977.4 2075.1 C -f -S -n -1952.2 2051.3 m -1953.6 2050.7 1955.1 2050.1 1956.5 2049.4 C -1955.1 2050.1 1953.6 2050.7 1952.2 2051.3 C -f -S -n -1952 2051.1 m -1951.2 2050.3 1950.3 2049.5 1949.6 2048.7 C -1950.3 2049.5 1951.2 2050.3 1952 2051.1 C -f -S -n -1946 2045.3 m -1947.3 2045.9 1948.1 2047 1949.1 2048.2 C -1948.6 2046.8 1947.1 2045.8 1946 2045.3 C -f -S -n -1937.3 2036 m -1937.4 2035.5 1937 2036.5 1937.6 2036.9 C -1938.8 2037.9 1940.1 2038.8 1940.9 2040.3 C -1940.6 2038.2 1937.6 2038.2 1937.3 2036 C -f -S -n -1935.2 2073.2 m -1936.4 2069.9 1935.8 2061.8 1935.6 2056.4 C -1935.8 2055.9 1936.3 2055.7 1936.1 2055.2 C -1935.7 2054.7 1935 2055 1934.4 2054.9 C -1934.4 2061.5 1934.4 2068.7 1934.4 2074.6 C -1935.7 2075.1 1936 2073.7 1935.2 2073.2 C -[0 0.01 1 0] vc -f -S -n -vmrs -1939 2030.7 m -1940.3 2027.4 1939.7 2019.3 1939.5 2013.9 C -1939.7 2013.5 1940.1 2013.2 1940 2012.7 C -1939.5 2012.3 1938.8 2012.5 1938.3 2012.4 C -1938.3 2019 1938.3 2026.2 1938.3 2032.1 C -1939.5 2032.7 1939.8 2031.2 1939 2030.7 C -[0 0.01 1 0] vc -f -0.4 w -2 J -2 M -S -n -1975.2 2077.2 m -1975.3 2077.3 1975.4 2077.4 1975.5 2077.5 C -1974.7 2073.2 1974.9 2067.5 1975.2 2063.6 C -1975.4 2064 1974.6 2063.9 1974.8 2064.3 C -1974.9 2069.9 1974.3 2076.5 1975.2 2081.1 C -1974.9 2079.9 1974.9 2078.4 1975.2 2077.2 C -[0.92 0.92 0 0.67] vc -f -S -n -1930.8 2067.4 m -1931.5 2070.1 1929.6 2072.1 1930.6 2074.6 C -1931 2072.6 1930.8 2069.8 1930.8 2067.4 C -f -S -n -2010 2050.1 m -2009.8 2050.5 2009.5 2050.9 2009.3 2051.1 C -2009.5 2056.7 2008.9 2063.3 2009.8 2067.9 C -2009.5 2062.1 2009.3 2054.7 2010 2050.1 C -f -S -n -1930.1 2060.9 m -1929.3 2057.1 1930.7 2054.8 1929.9 2051.3 C -1930.2 2050.2 1931.1 2049.6 1931.8 2049.2 C -1931.4 2049.6 1930.4 2049.5 1930.1 2050.1 C -1928.4 2054.8 1933.4 2063.5 1925.3 2064.3 C -1927.2 2063.9 1928.5 2062.1 1930.1 2060.9 C -[0.07 0.06 0 0.58] vc -f -S -n -1929.6 2061.2 m -1929.6 2057.6 1929.6 2054.1 1929.6 2050.6 C -1930 2049.9 1930.5 2049.4 1931.1 2049.2 C -1930 2048.6 1930.5 2050.2 1929.4 2049.6 C -1928 2054.4 1932.8 2063 1925.3 2064 C -1926.9 2063.3 1928.3 2062.4 1929.6 2061.2 C -[0.4 0.4 0 0] vc -f -S -n -1930.8 2061.6 m -1930.5 2058 1931.6 2054 1930.8 2051.3 C -1930.3 2054.5 1930.9 2058.5 1930.4 2061.9 C -1930.5 2061.2 1931 2062.2 1930.8 2061.6 C -[0.92 0.92 0 0.67] vc -f -S -n -1941.2 2045.1 m -1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C -1934.2 2040 1933.7 2036.4 1934 2039.3 C -1934.9 2040.1 1936.1 2039.9 1936.8 2040.8 C -1935.3 2044.2 1942.3 2041.7 1939.5 2046 C -1937.1 2048.5 1940.5 2045.6 1941.2 2045.1 C -f -S -n -1910 2045.8 m -1910 2039.4 1910 2033 1910 2026.6 C -1910 2033 1910 2039.4 1910 2045.8 C -f -S -n -1978.8 2022.3 m -1979.1 2021.7 1979.4 2020.4 1978.6 2021.6 C -1978.6 2026.9 1978.6 2033 1978.6 2037.6 C -1979.2 2037 1979.1 2038.2 1979.1 2038.6 C -1978.7 2033.6 1978.9 2026.8 1978.8 2022.3 C -f -S -n -vmrs -2026.1 2041.2 m -2026.1 2034.8 2026.1 2028.3 2026.1 2021.8 C -2026.1 2028.5 2026.3 2035.4 2025.9 2042 C -2024.4 2042.9 2022.9 2044.1 2021.3 2044.8 C -2023.1 2044 2025.1 2042.8 2026.1 2041.2 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -2026.4 2021.8 m -2026.3 2028.5 2026.5 2035.4 2026.1 2042 C -2025.6 2042.8 2024.7 2042.7 2024.2 2043.4 C -2024.7 2042.7 2025.5 2042.7 2026.1 2042.2 C -2026.5 2035.5 2026.3 2027.9 2026.4 2021.8 C -[0.4 0.4 0 0] vc -f -S -n -2025.6 2038.4 m -2025.6 2033 2025.6 2027.6 2025.6 2022.3 C -2025.6 2027.6 2025.6 2033 2025.6 2038.4 C -[0.92 0.92 0 0.67] vc -f -S -n -1934 2023.5 m -1934 2024.7 1933.8 2026 1934.2 2027.1 C -1934 2025.5 1934.7 2024.6 1934 2023.5 C -f -S -n -1928.2 2023.5 m -1928 2024.6 1927.4 2023.1 1926.8 2023.2 C -1926.2 2021 1921.4 2019.3 1923.2 2018 C -1922.7 2016.5 1923.2 2019.3 1922.2 2018.2 C -1924.4 2020.4 1926.2 2023.3 1928.9 2024.9 C -1927.9 2024.2 1929.8 2023.5 1928.2 2023.5 C -[0.18 0.18 0 0.78] vc -f -S -n -1934 2019.2 m -1932 2019.6 1930.8 2022.6 1928.7 2021.8 C -1924.5 2016.5 1918.2 2011.8 1914 2006.7 C -1914 2005.7 1914 2004.6 1914 2003.6 C -1913.6 2004.3 1913.9 2005.8 1913.8 2006.9 C -1919 2012.4 1924.1 2016.5 1929.2 2022.3 C -1931 2021.7 1932.2 2019.8 1934 2019.2 C -f -S -n -1928.7 2024.9 m -1926.3 2022.7 1924.1 2020.4 1921.7 2018.2 C -1924.1 2020.4 1926.3 2022.7 1928.7 2024.9 C -[0.65 0.65 0 0.42] vc -f -S -n -1914.3 2006.7 m -1918.7 2011.8 1924.5 2016.4 1928.9 2021.6 C -1924.2 2016.1 1919 2012.1 1914.3 2006.7 C -[0.07 0.06 0 0.58] vc -f -S -n -1924.8 2020.8 m -1921.2 2016.9 1925.6 2022.5 1926 2021.1 C -1924.2 2021 1926.7 2019.6 1924.8 2020.8 C -[0.92 0.92 0 0.67] vc -f -S -n -1934 2018.4 m -1933.2 2014.7 1934.5 2012.3 1933.7 2008.8 C -1934 2007.8 1935 2007.2 1935.6 2006.7 C -1935.3 2007.1 1934.3 2007 1934 2007.6 C -1932.2 2012.3 1937.2 2021 1929.2 2021.8 C -1931.1 2021.4 1932.3 2019.6 1934 2018.4 C -[0.07 0.06 0 0.58] vc -f -S -n -vmrs -1933.5 2018.7 m -1933.5 2015.1 1933.5 2011.7 1933.5 2008.1 C -1933.8 2007.4 1934.3 2006.9 1934.9 2006.7 C -1933.8 2006.1 1934.3 2007.7 1933.2 2007.2 C -1931.9 2012 1936.7 2020.5 1929.2 2021.6 C -1930.7 2020.8 1932.2 2019.9 1933.5 2018.7 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1934.7 2019.2 m -1934.3 2015.6 1935.4 2011.5 1934.7 2008.8 C -1934.1 2012 1934.7 2016 1934.2 2019.4 C -1934.4 2018.7 1934.8 2019.8 1934.7 2019.2 C -[0.92 0.92 0 0.67] vc -f -S -n -1917.6 2013.6 m -1917.8 2011.1 1916.8 2014.2 1917.2 2012.2 C -1916.3 2012.9 1914.8 2011.8 1914.3 2010.8 C -1914.2 2010.5 1914.4 2010.4 1914.5 2010.3 C -1913.9 2008.8 1913.9 2011.9 1914.3 2012 C -1916.3 2012 1917.6 2013.6 1916.7 2015.6 C -1913.7 2017.4 1919.6 2014.8 1917.6 2013.6 C -f -S -n -1887.2 2015.3 m -1887.2 2008.9 1887.2 2002.5 1887.2 1996.1 C -1887.2 2002.5 1887.2 2008.9 1887.2 2015.3 C -f -S -n -1916.7 2014.4 m -1917 2012.1 1913 2013 1913.8 2010.8 C -1912.1 2009.8 1910.9 2009.4 1910.7 2007.9 C -1910.4 2010.6 1913.4 2010.4 1914 2012.4 C -1914.9 2012.8 1916.6 2012.9 1916.4 2014.4 C -1916.9 2015.1 1914.5 2016.6 1916.2 2015.8 C -1916.4 2015.3 1916.7 2015 1916.7 2014.4 C -[0.65 0.65 0 0.42] vc -f -S -n -1914 2009.3 m -1912.8 2010.9 1909.6 2005.3 1911.9 2009.8 C -1912.3 2009.6 1913.6 2010.2 1914 2009.3 C -[0.92 0.92 0 0.67] vc -f -S -n -1951.2 1998.8 m -1949 1996.4 1951.5 1994 1950.3 1991.8 C -1949.1 1989.1 1954 1982.7 1948.8 1981.2 C -1949.2 1981.5 1951 1982.4 1950.8 1983.6 C -1951.9 1988.6 1947.1 1986.5 1948.1 1990.4 C -1948.5 1990.3 1948.7 1990.7 1948.6 1991.1 C -1949 1992.5 1947.3 1991.9 1948.1 1992.5 C -1947.1 1992.7 1945.7 1993.5 1945.2 1994.7 C -1944.5 1996.8 1947.7 2000.5 1943.8 2001.4 C -1943.4 2002 1943.7 2004 1942.4 2004.5 C -1945.2 2002.2 1948.9 2000.9 1951.2 1998.8 C -f -S -n -1994.9 1993 m -1995.1 1996.5 1994.5 2000.3 1995.4 2003.6 C -1994.5 2000.3 1995.1 1996.5 1994.9 1993 C -f -S -n -1913.8 2003.3 m -1913.8 1996.9 1913.8 1990.5 1913.8 1984.1 C -1913.8 1990.5 1913.8 1996.9 1913.8 2003.3 C -f -S -n -1941.9 1998 m -1940.5 1997.3 1940.7 1999.4 1940.7 2000 C -1942.8 2001.3 1942.6 1998.8 1941.9 1998 C -[0 0 0 0] vc -f -S -n -vmrs -1942.1 1999.2 m -1942.2 1998.9 1941.8 1998.8 1941.6 1998.5 C -1940.4 1998 1940.7 1999.7 1940.7 2000 C -1941.6 2000.3 1942.6 2000.4 1942.1 1999.2 C -[0.92 0.92 0 0.67] vc -f -0.4 w -2 J -2 M -S -n -1940 1997.1 m -1939.8 1996 1939.7 1995.9 1939.2 1995.2 C -1939.1 1995.3 1938.5 1997.9 1937.8 1996.4 C -1938 1997.3 1939.4 1998.6 1940 1997.1 C -f -S -n -1911.2 1995.9 m -1911.2 1991.6 1911.3 1987.2 1911.4 1982.9 C -1911.3 1987.2 1911.2 1991.6 1911.2 1995.9 C -f -S -n -1947.2 1979.1 m -1945.1 1978.8 1944.6 1975.7 1942.4 1975 C -1940.5 1972.6 1942.2 1973.7 1942.4 1975.7 C -1945.8 1975.5 1944.2 1979.8 1947.6 1979.6 C -1948.3 1982.3 1948.5 1980 1947.2 1979.1 C -f -S -n -1939.5 1973.3 m -1940.1 1972.6 1939.8 1974.2 1940.2 1973.1 C -1939.1 1972.8 1938.8 1968.5 1935.9 1969.7 C -1937.4 1969.2 1938.5 1970.6 1939 1971.4 C -1939.2 1972.7 1938.6 1973.9 1939.5 1973.3 C -f -S -n -1975.2 2073.2 m -1975.2 2070.2 1975.2 2067.2 1975.2 2064.3 C -1975.2 2067.2 1975.2 2070.2 1975.2 2073.2 C -[0.18 0.18 0 0.78] vc -f -S -n -1929.9 2065.7 m -1928.1 2065.6 1926 2068.8 1924.1 2066.9 C -1918.1 2060.9 1912.9 2055.7 1907.1 2049.9 C -1906.7 2047.1 1906.9 2043.9 1906.8 2041 C -1906.8 2043.9 1906.8 2046.8 1906.8 2049.6 C -1913.2 2055.5 1918.7 2061.9 1925.1 2067.6 C -1927.1 2067.9 1928.6 2064.4 1930.1 2066.2 C -1929.7 2070.3 1929.9 2074.7 1929.9 2078.9 C -1929.6 2074.4 1930.5 2070.1 1929.9 2065.7 C -[0.07 0.06 0 0.58] vc -f -S -n -1930.1 2061.6 m -1928.1 2062.1 1927 2065.1 1924.8 2064.3 C -1920.7 2058.9 1914.4 2054.3 1910.2 2049.2 C -1910.2 2048.1 1910.2 2047.1 1910.2 2046 C -1909.8 2046.8 1910 2048.3 1910 2049.4 C -1915.1 2054.9 1920.3 2059 1925.3 2064.8 C -1927.1 2064.2 1928.4 2062.3 1930.1 2061.6 C -[0.18 0.18 0 0.78] vc -f -S -n -1932 2049.9 m -1932.3 2050.3 1932 2050.4 1932.8 2050.4 C -1932 2050.4 1932.2 2049.2 1931.3 2049.6 C -1931.4 2050.5 1930.3 2050.4 1930.4 2051.3 C -1931.1 2051.1 1930.7 2049.4 1932 2049.9 C -f -S -n -1938.3 2046 m -1936.3 2046.8 1935.2 2047.2 1934.2 2048.9 C -1935.3 2047.7 1936.8 2046.2 1938.3 2046 C -[0.4 0.4 0 0] vc -f -S -n -vmrs -1938.3 2047 m -1937.9 2046.9 1936.6 2047.1 1936.1 2048 C -1936.5 2047.5 1937.3 2046.7 1938.3 2047 C -[0.18 0.18 0 0.78] vc -f -0.4 w -2 J -2 M -S -n -1910.2 2043.2 m -1910.1 2037.5 1910 2031.8 1910 2026.1 C -1910 2031.8 1910.1 2037.5 1910.2 2043.2 C -f -S -n -1933.5 2032.1 m -1933.7 2035.2 1932.8 2035.8 1933.7 2038.6 C -1933.3 2036.6 1934.6 2018 1933.5 2032.1 C -f -S -n -1907.3 2021.8 m -1906.6 2025.9 1909.4 2032.6 1903.2 2034 C -1902.8 2034.1 1902.4 2033.9 1902 2033.8 C -1897.9 2028.5 1891.6 2023.8 1887.4 2018.7 C -1887.4 2017.7 1887.4 2016.6 1887.4 2015.6 C -1887 2016.3 1887.2 2017.8 1887.2 2018.9 C -1892.3 2024.4 1897.5 2028.5 1902.5 2034.3 C -1904.3 2033.6 1905.7 2032 1907.3 2030.9 C -1907.3 2027.9 1907.3 2024.9 1907.3 2021.8 C -f -S -n -1933.7 2023.2 m -1932 2021.7 1931.1 2024.9 1929.4 2024.9 C -1931.2 2024.7 1932.4 2021.5 1933.7 2023.2 C -f -S -n -1989.2 2024.4 m -1987.4 2023.7 1985.8 2022.2 1985.1 2020.4 C -1984.6 2020.1 1986 2018.9 1985.1 2019.2 C -1985.6 2020.8 1984.1 2019.4 1984.6 2021.1 C -1986.3 2022.3 1988.1 2025.3 1989.2 2024.4 C -f -S -n -1904.4 2031.9 m -1903 2029.7 1905.3 2027.7 1904.2 2025.9 C -1904.5 2025 1903.7 2023 1904 2021.3 C -1904 2022.3 1903.2 2022 1902.5 2022 C -1901.3 2022.3 1902.2 2020.1 1901.6 2019.6 C -1902.5 2019.8 1902.6 2018.3 1903.5 2018.9 C -1903.7 2021.8 1905.6 2016.8 1905.6 2020.6 C -1905.9 2020 1906.3 2020.8 1906.1 2021.1 C -1905.8 2022.7 1906.7 2020.4 1906.4 2019.9 C -1906.4 2018.5 1908.2 2017.8 1906.8 2016.5 C -1906.9 2015.7 1907.7 2017.1 1907.1 2016.3 C -1908.5 2015.8 1910.3 2015.1 1911.6 2016 C -1912.2 2016.2 1911.9 2018 1911.6 2018 C -1914.5 2017.1 1910.4 2013.6 1913.3 2013.4 C -1912.4 2011.3 1910.5 2011.8 1909.5 2010 C -1910 2010.5 1909 2010.8 1908.8 2011.2 C -1907.5 2009.9 1906.1 2011.7 1904.9 2011.5 C -1904.7 2010.9 1904.3 2010.5 1904.4 2009.8 C -1905 2010.2 1904.6 2008.6 1905.4 2008.1 C -1906.6 2007.5 1907.7 2008.4 1908.5 2007.4 C -1908.9 2008.5 1909.7 2008.1 1909 2007.2 C -1908.1 2006.5 1905.9 2007.3 1905.4 2007.4 C -1903.9 2007.3 1905.2 2008.5 1904.2 2008.4 C -1904.6 2009.9 1902.8 2010.3 1902.3 2010.5 C -1901.5 2009.9 1900.4 2010 1899.4 2010 C -1898.6 2011.2 1898.2 2013.4 1896.5 2013.4 C -1896 2012.9 1894.4 2012.9 1893.6 2012.9 C -1893.1 2013.9 1892.9 2015.5 1891.5 2016 C -1890.3 2016.1 1889.2 2014 1888.6 2015.8 C -1890 2016 1891 2016.9 1892.9 2016.5 C -1894.1 2017.2 1892.8 2018.3 1893.2 2018.9 C -1892.6 2018.9 1891.1 2019.8 1890.5 2020.6 C -1891.1 2023.6 1893.2 2019.8 1893.9 2022.5 C -1894.1 2023.3 1892.7 2023.6 1893.9 2024 C -1894.2 2024.3 1897.4 2023.8 1896.5 2026.1 C -1896 2025.6 1897.4 2028.1 1897.5 2027.1 C -1898.4 2027.4 1899.3 2027 1899.6 2028.5 C -1899.5 2028.6 1899.4 2028.8 1899.2 2028.8 C -1899.3 2029.2 1899.6 2029.8 1900.1 2030.2 C -1900.4 2029.6 1901 2030 1901.8 2030.2 C -1903.1 2032.1 1900.4 2031.5 1902.8 2033.1 C -1903.3 2032.7 1904.5 2032 1904.4 2031.9 C -[0.21 0.21 0 0] vc -f -S -n -1909.2 2019.4 m -1908.8 2020.3 1910.2 2019.8 1909.2 2019.2 C -1908.3 2019.3 1907.6 2020.2 1907.6 2021.3 C -1908.5 2021 1907.6 2019 1909.2 2019.4 C -[0.18 0.18 0 0.78] vc -f -S -n -1915.5 2015.6 m -1913.5 2016.3 1912.4 2016.8 1911.4 2018.4 C -1912.5 2017.2 1914 2015.7 1915.5 2015.6 C -[0.4 0.4 0 0] vc -f -S -n -1915.5 2016.5 m -1915.1 2016.4 1913.8 2016.6 1913.3 2017.5 C -1913.7 2017 1914.5 2016.2 1915.5 2016.5 C -[0.18 0.18 0 0.78] vc -f -S -n -vmrs -1887.4 2012.7 m -1887.3 2007 1887.2 2001.3 1887.2 1995.6 C -1887.2 2001.3 1887.3 2007 1887.4 2012.7 C -[0.18 0.18 0 0.78] vc -f -0.4 w -2 J -2 M -S -n -1935.9 2007.4 m -1936.2 2007.8 1935.8 2007.9 1936.6 2007.9 C -1935.9 2007.9 1936.1 2006.7 1935.2 2007.2 C -1935.2 2008.1 1934.1 2007.9 1934.2 2008.8 C -1935 2008.7 1934.6 2006.9 1935.9 2007.4 C -f -S -n -1942.1 2003.6 m -1940.1 2004.3 1939.1 2004.8 1938 2006.4 C -1939.1 2005.2 1940.6 2003.7 1942.1 2003.6 C -[0.4 0.4 0 0] vc -f -S -n -1942.1 2004.5 m -1941.8 2004.4 1940.4 2004.6 1940 2005.5 C -1940.4 2005 1941.2 2004.2 1942.1 2004.5 C -[0.18 0.18 0 0.78] vc -f -S -n -1914 2000.7 m -1914 1995 1913.9 1989.3 1913.8 1983.6 C -1913.9 1989.3 1914 1995 1914 2000.7 C -f -S -n -1941.6 1998.3 m -1943.4 2001.9 1942.4 1996 1940.9 1998.3 C -1941.2 1998.3 1941.4 1998.3 1941.6 1998.3 C -f -S -n -1954.8 1989.9 m -1953.9 1989.6 1954.7 1991.6 1953.9 1991.1 C -1954.5 1993.1 1953.6 1998 1954.6 1993.2 C -1954 1992.2 1954.7 1990.7 1954.8 1989.9 C -f -S -n -1947.6 1992.5 m -1946.2 1993.5 1944.9 1993 1944.8 1994.7 C -1945.5 1994 1947 1992.2 1947.6 1992.5 C -f -S -n -1910.7 1982.2 m -1910.3 1981.8 1909.7 1982 1909.2 1982 C -1909.7 1982 1910.3 1981.9 1910.7 1982.2 C -1911 1987.1 1910 1992.6 1910.7 1997.3 C -1910.7 1992.3 1910.7 1987.2 1910.7 1982.2 C -[0.65 0.65 0 0.42] vc -f -S -n -1910.9 1992.8 m -1910.9 1991.3 1910.9 1989.7 1910.9 1988.2 C -1910.9 1989.7 1910.9 1991.3 1910.9 1992.8 C -[0.18 0.18 0 0.78] vc -f -S -n -vmrs -1953.6 1983.6 m -1954.1 1985.3 1953.2 1988.6 1954.8 1989.4 C -1954.1 1987.9 1954.4 1985.4 1953.6 1983.6 C -[0.18 0.18 0 0.78] vc -f -0.4 w -2 J -2 M -S -n -1910.7 1982 m -1911.6 1982.9 1911 1984.4 1911.2 1985.6 C -1911 1984.4 1911.6 1982.9 1910.7 1982 C -f -S -n -1947.2 1979.6 m -1947.5 1980.6 1948.3 1980.6 1947.4 1979.6 C -1946.2 1979.4 1945.7 1978.8 1947.2 1979.6 C -f -S -n -1930.4 2061.4 m -1930.4 2058 1930.4 2053.5 1930.4 2051.1 C -1930.7 2054.6 1929.8 2057.4 1930.1 2061.2 C -1929.5 2061.9 1929.7 2061.2 1930.4 2061.4 C -[0.65 0.65 0 0.42] vc -f -S -n -1939.5 2044.8 m -1940 2041.5 1935.2 2044.3 1936.4 2040.8 C -1934.9 2040.9 1934.1 2039.7 1933.5 2038.6 C -1933.3 2035.4 1933.2 2040 1934 2040.3 C -1936.2 2040.6 1936.3 2043.6 1938.5 2043.4 C -1939.7 2044.2 1939.4 2045.6 1938.3 2046.5 C -1939.1 2046.6 1939.6 2045.6 1939.5 2044.8 C -f -S -n -1910.4 2045.3 m -1910.4 2039.5 1910.4 2033.6 1910.4 2027.8 C -1910.4 2033.6 1910.4 2039.5 1910.4 2045.3 C -f -S -n -1906.8 2030.9 m -1907.6 2026.8 1905 2020.8 1909 2018.7 C -1906.5 2018.9 1906.8 2022.4 1906.8 2024.7 C -1906.4 2028.2 1907.9 2032 1903 2033.8 C -1902.2 2034 1903.8 2033.4 1904.2 2033.1 C -1905.1 2032.4 1905.9 2031.5 1906.8 2030.9 C -[0.07 0.06 0 0.58] vc -f -S -n -1907.1 2030.7 m -1907.1 2028.8 1907.1 2027 1907.1 2025.2 C -1907.1 2027 1907.1 2028.8 1907.1 2030.7 C -[0.65 0.65 0 0.42] vc -f -S -n -1932 2023.2 m -1932.2 2023.6 1931.7 2023.7 1931.6 2024 C -1932 2023.7 1932.3 2022.8 1933 2023 C -1933.9 2024.3 1933.3 2026.2 1933.5 2027.8 C -1933.5 2026.4 1934.9 2022.2 1932 2023.2 C -f -S -n -2026.1 2021.6 m -2026.1 2020.8 2026.1 2019.9 2026.1 2019.2 C -2026.1 2019.9 2026.1 2020.8 2026.1 2021.6 C -f -S -n -vmrs -1934.2 2018.9 m -1934.2 2015.5 1934.2 2011 1934.2 2008.6 C -1934.5 2012.1 1933.7 2014.9 1934 2018.7 C -1933.4 2019.5 1933.5 2018.7 1934.2 2018.9 C -[0.65 0.65 0 0.42] vc -f -0.4 w -2 J -2 M -S -n -1887.6 2014.8 m -1887.6 2009 1887.6 2003.1 1887.6 1997.3 C -1887.6 2003.1 1887.6 2009 1887.6 2014.8 C -f -S -n -1914.3 2002.8 m -1914.3 1997 1914.3 1991.1 1914.3 1985.3 C -1914.3 1991.1 1914.3 1997 1914.3 2002.8 C -f -S -n -1995.4 1992.3 m -1995.4 1991.5 1995.4 1990.7 1995.4 1989.9 C -1995.4 1990.7 1995.4 1991.5 1995.4 1992.3 C -f -S -n -1896 1988.4 m -1896.9 1988 1897.8 1987.7 1898.7 1987.2 C -1897.8 1987.7 1896.9 1988 1896 1988.4 C -f -S -n -1899.4 1986.8 m -1900.4 1986.3 1901.3 1985.8 1902.3 1985.3 C -1901.3 1985.8 1900.4 1986.3 1899.4 1986.8 C -f -S -n -1902.8 1985.1 m -1905.2 1984 1905.2 1984 1902.8 1985.1 C -f -S -n -1949.1 1983.4 m -1950.2 1984.4 1947.8 1984.6 1949.3 1985.1 C -1949.5 1984.4 1949.6 1984.1 1949.1 1983.4 C -[0.07 0.06 0 0.58] vc -f -S -n -1906.1 1983.4 m -1908.6 1982 1908.6 1982 1906.1 1983.4 C -[0.65 0.65 0 0.42] vc -f -S -n -1922.7 1976.4 m -1923.6 1976 1924.4 1975.7 1925.3 1975.2 C -1924.4 1975.7 1923.6 1976 1922.7 1976.4 C -f -S -n -vmrs -1926 1974.8 m -1927 1974.3 1928 1973.8 1928.9 1973.3 C -1928 1973.8 1927 1974.3 1926 1974.8 C -[0.65 0.65 0 0.42] vc -f -0.4 w -2 J -2 M -S -n -1929.4 1973.1 m -1931.9 1972 1931.9 1972 1929.4 1973.1 C -f -S -n -1932.8 1971.4 m -1935.3 1970 1935.3 1970 1932.8 1971.4 C -f -S -n -1949.6 2097.2 m -1951.1 2096.4 1952.6 2095.5 1954.1 2094.8 C -1952.6 2095.5 1951.1 2096.4 1949.6 2097.2 C -[0.07 0.06 0 0.58] vc -f -S -n -1955.1 2094.3 m -1956.7 2093.5 1958.3 2092.7 1959.9 2091.9 C -1958.3 2092.7 1956.7 2093.5 1955.1 2094.3 C -f -S -n -1960.4 2091.6 m -1961.3 2091.2 1962.1 2090.9 1963 2090.4 C -1962.1 2090.9 1961.3 2091.2 1960.4 2091.6 C -f -S -n -1963.5 2090.2 m -1964.4 2089.7 1965.2 2089.2 1966.1 2088.8 C -1965.2 2089.2 1964.4 2089.7 1963.5 2090.2 C -f -S -n -1966.6 2088.5 m -1969.5 2087.1 1972.4 2085.8 1975.2 2084.4 C -1972.4 2085.8 1969.5 2087.1 1966.6 2088.5 C -f -S -n -1965.2 2086.1 m -1965.9 2085.7 1966.8 2085.3 1967.6 2084.9 C -1966.8 2085.3 1965.9 2085.7 1965.2 2086.1 C -f -S -n -1968.3 2084.7 m -1969.2 2084.3 1970 2083.9 1970.9 2083.5 C -1970 2083.9 1969.2 2084.3 1968.3 2084.7 C -f -S -n -vmrs -1984.1 2084 m -1985.6 2083.2 1987.2 2082.3 1988.7 2081.6 C -1987.2 2082.3 1985.6 2083.2 1984.1 2084 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1976 2078.7 m -1978.1 2080.1 1980 2082 1982 2083.7 C -1980 2081.9 1977.9 2080.3 1976 2078.2 C -1975.5 2079.9 1975.8 2081.9 1975.7 2083.7 C -1975.8 2082 1975.5 2080.2 1976 2078.7 C -f -S -n -1989.6 2081.1 m -1991.3 2080.3 1992.8 2079.5 1994.4 2078.7 C -1992.8 2079.5 1991.3 2080.3 1989.6 2081.1 C -f -S -n -1933.2 2074.6 m -1932.4 2076.2 1932.8 2077.5 1933 2078.7 C -1933 2077.6 1932.9 2074.8 1933.2 2074.6 C -f -S -n -1994.9 2078.4 m -1995.8 2078 1996.7 2077.7 1997.6 2077.2 C -1996.7 2077.7 1995.8 2078 1994.9 2078.4 C -f -S -n -1998 2077 m -1998.9 2076.5 1999.8 2076 2000.7 2075.6 C -1999.8 2076 1998.9 2076.5 1998 2077 C -f -S -n -2001.2 2075.3 m -2004 2073.9 2006.9 2072.6 2009.8 2071.2 C -2006.9 2072.6 2004 2073.9 2001.2 2075.3 C -f -S -n -1980.5 2060.7 m -1979.9 2060.7 1976.7 2062.8 1975.7 2064.5 C -1975.7 2067.5 1975.7 2070.5 1975.7 2073.4 C -1976.3 2068.7 1973.9 2061.6 1980.5 2060.7 C -f -S -n -1999.7 2072.9 m -2000.5 2072.5 2001.3 2072.1 2002.1 2071.7 C -2001.3 2072.1 2000.5 2072.5 1999.7 2072.9 C -f -S -n -2002.8 2071.5 m -2003.7 2071.1 2004.6 2070.7 2005.5 2070.3 C -2004.6 2070.7 2003.7 2071.1 2002.8 2071.5 C -f -S -n -vmrs -2015.1 2047.5 m -2014.4 2047.5 2011.2 2049.6 2010.3 2051.3 C -2010.3 2057.7 2010.3 2064.1 2010.3 2070.5 C -2010.3 2063.9 2010.1 2057.1 2010.5 2050.6 C -2012 2049.3 2013.5 2048.3 2015.1 2047.5 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1910.4 2049.2 m -1914.8 2054.3 1920.7 2058.9 1925.1 2064 C -1920.4 2058.6 1915.1 2054.6 1910.4 2049.2 C -f -S -n -1988.2 2057.3 m -1989.1 2056.8 1989.9 2056.2 1990.8 2055.6 C -1989.9 2056.2 1989.1 2056.8 1988.2 2057.3 C -f -S -n -1991.6 2051.3 m -1991.6 2046.3 1991.6 2041.2 1991.6 2036.2 C -1991.6 2041.2 1991.6 2046.3 1991.6 2051.3 C -f -S -n -1935.6 2047.5 m -1932.9 2051.7 1939.7 2043.8 1935.6 2047.5 C -f -S -n -1938.8 2043.9 m -1938.1 2043.3 1938.2 2043.7 1937.3 2043.4 C -1938.7 2043 1938.2 2044.9 1939 2045.3 C -1938.2 2045.3 1938.7 2046.6 1937.8 2046.5 C -1939.1 2046.2 1939.1 2044.5 1938.8 2043.9 C -f -S -n -1972.4 2045.6 m -1973.4 2045 1974.5 2044.4 1975.5 2043.9 C -1974.5 2044.4 1973.4 2045 1972.4 2045.6 C -f -S -n -1969 2043.6 m -1969.8 2043.2 1970.6 2042.9 1971.4 2042.4 C -1970.6 2042.9 1969.8 2043.2 1969 2043.6 C -f -S -n -1972.1 2042.2 m -1973 2041.8 1973.9 2041.4 1974.8 2041 C -1973.9 2041.4 1973 2041.8 1972.1 2042.2 C -f -S -n -1906.6 2035 m -1905 2034.7 1904.8 2036.6 1903.5 2036.9 C -1904.9 2037 1905.8 2033.4 1907.1 2035.7 C -1907.1 2037.2 1907.1 2038.6 1907.1 2040 C -1906.9 2038.4 1907.5 2036.4 1906.6 2035 C -f -S -n -vmrs -1937.1 2032.1 m -1936.2 2033.7 1936.6 2035 1936.8 2036.2 C -1936.8 2035.1 1936.8 2032.4 1937.1 2032.1 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1887.6 2018.7 m -1892 2023.8 1897.9 2028.4 1902.3 2033.6 C -1897.6 2028.1 1892.3 2024.1 1887.6 2018.7 C -f -S -n -1999.7 2031.4 m -1998.7 2030.3 1997.6 2029.2 1996.6 2028 C -1997.6 2029.2 1998.7 2030.3 1999.7 2031.4 C -f -S -n -1912.8 2017 m -1910.6 2021.1 1913.6 2015.3 1914.5 2016 C -1914 2016.3 1913.4 2016.7 1912.8 2017 C -f -S -n -1939.5 2005 m -1936.7 2009.2 1943.6 2001.3 1939.5 2005 C -f -S -n -1942.6 2001.4 m -1941.9 2000.8 1942 2001.2 1941.2 2000.9 C -1942.5 2000.6 1942.1 2002.4 1942.8 2002.8 C -1942 2002.8 1942.5 2004.1 1941.6 2004 C -1943 2003.7 1942.9 2002.1 1942.6 2001.4 C -f -S -n -2006.2 2000.7 m -2005.4 2001.5 2004 2002.8 2004 2002.8 C -2004.5 2002.4 2005.5 2001.4 2006.2 2000.7 C -f -S -n -1998.5 2001.6 m -1997.7 2002 1996.8 2002.4 1995.9 2002.6 C -1995.5 1999.3 1995.7 1995.7 1995.6 1992.3 C -1995.6 1995.7 1995.6 1999.2 1995.6 2002.6 C -1996.6 2002.4 1997.7 2002.2 1998.5 2001.6 C -[0.4 0.4 0 0] vc -f -S -n -1996.1 2002.8 m -1995.9 2002.8 1995.8 2002.8 1995.6 2002.8 C -1995.2 1999.5 1995.5 1995.9 1995.4 1992.5 C -1995.4 1995.9 1995.4 1999.4 1995.4 2002.8 C -1996.4 2003.1 1998.2 2001.6 1996.1 2002.8 C -[0.07 0.06 0 0.58] vc -f -S -n -1969 2002.1 m -1968 2001 1966.9 1999.9 1965.9 1998.8 C -1966.9 1999.9 1968 2001 1969 2002.1 C -f -S -n -vmrs -2000 2001.2 m -2002.1 2000 2004.1 1998.9 2006.2 1997.8 C -2004.1 1998.9 2002.1 2000 2000 2001.2 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1895.8 1984.8 m -1898.3 1983.6 1900.8 1982.3 1903.2 1981 C -1900.8 1982.3 1898.3 1983.6 1895.8 1984.8 C -f -S -n -1905.2 1980.3 m -1906.4 1979.9 1907.6 1979.5 1908.8 1979.1 C -1907.6 1979.5 1906.4 1979.9 1905.2 1980.3 C -f -S -n -1964.7 1977.4 m -1963.8 1977.5 1962.5 1980.2 1960.8 1980 C -1962.5 1980.2 1963.3 1978 1964.7 1977.4 C -f -S -n -1952 1979.6 m -1955.2 1979.2 1955.2 1979.2 1952 1979.6 C -f -S -n -1937.8 1966.4 m -1941.2 1969.5 1946.1 1976.4 1951.5 1979.3 C -1946.1 1976.7 1942.8 1970.4 1937.8 1966.4 C -f -S -n -1911.9 1978.6 m -1914.3 1977.4 1916.7 1976.2 1919.1 1975 C -1916.7 1976.2 1914.3 1977.4 1911.9 1978.6 C -f -S -n -1975.5 1971.4 m -1974.6 1972.2 1973.3 1973.6 1973.3 1973.6 C -1973.7 1973.1 1974.8 1972.1 1975.5 1971.4 C -f -S -n -1922.4 1972.8 m -1924.9 1971.6 1927.4 1970.3 1929.9 1969 C -1927.4 1970.3 1924.9 1971.6 1922.4 1972.8 C -f -S -n -1969.2 1971.9 m -1971.1 1970.9 1972.9 1969.8 1974.8 1968.8 C -1972.9 1969.8 1971.1 1970.9 1969.2 1971.9 C -f -S -n -vmrs -1931.8 1968.3 m -1933 1967.9 1934.2 1967.5 1935.4 1967.1 C -1934.2 1967.5 1933 1967.9 1931.8 1968.3 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1940.7 2072.4 m -1941.5 2072.4 1942.3 2072.3 1943.1 2072.2 C -1942.3 2072.3 1941.5 2072.4 1940.7 2072.4 C -[0 0 0 0.18] vc -f -S -n -1948.6 2069.3 m -1947 2069.5 1945.7 2068.9 1944.8 2069.8 C -1945.9 2068.5 1948.4 2070.2 1948.6 2069.3 C -f -S -n -1954.6 2066.4 m -1954.7 2067.9 1955.6 2067.3 1955.6 2068.8 C -1955.4 2067.8 1956 2066.6 1954.6 2066.4 C -f -S -n -1929.2 2061.2 m -1927.8 2062.1 1926.3 2064.1 1924.8 2063.3 C -1926.3 2064.6 1928 2062 1929.2 2061.2 C -f -S -n -1924.4 2067.4 m -1918.5 2061.6 1912.7 2055.9 1906.8 2050.1 C -1912.7 2055.9 1918.5 2061.6 1924.4 2067.4 C -[0.4 0.4 0 0] vc -f -S -n -1924.6 2062.8 m -1923.9 2062.1 1923.2 2061.2 1922.4 2060.4 C -1923.2 2061.2 1923.9 2062.1 1924.6 2062.8 C -[0 0 0 0.18] vc -f -S -n -1919.3 2057.3 m -1917.5 2055.6 1915.7 2053.8 1913.8 2052 C -1915.7 2053.8 1917.5 2055.6 1919.3 2057.3 C -f -S -n -1929.2 2055.2 m -1929.2 2054.2 1929.2 2053.2 1929.2 2052.3 C -1929.2 2053.2 1929.2 2054.2 1929.2 2055.2 C -f -S -n -1926.3 2049.6 m -1925.4 2049 1925.4 2050.5 1924.4 2050.4 C -1925.3 2051.3 1924.5 2051.9 1925.6 2052.5 C -1926.9 2052.6 1926 2050.6 1926.3 2049.6 C -f -S -n -vmrs -1911.2 2046.8 m -1910.1 2048.9 1911.9 2050.1 1913.1 2051.3 C -1912.1 2049.9 1910.6 2048.8 1911.2 2046.8 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1934 2048.7 m -1932.6 2048.7 1930.1 2047.7 1929.6 2049.4 C -1930.9 2048.6 1933.3 2049 1934 2048.7 C -f -S -n -1980 2048.4 m -1979.5 2046.8 1976.3 2047.9 1977.2 2045.6 C -1976.8 2045.1 1976.1 2044.7 1975.2 2044.8 C -1973.7 2046 1976.3 2046.4 1976.7 2047.5 C -1977.8 2047.2 1978.2 2050 1979.6 2049.2 C -1980 2049 1979.6 2048.6 1980 2048.4 C -f -S -n -1938.3 2045.6 m -1938.2 2044.4 1936.8 2043.8 1935.9 2043.4 C -1936.4 2044.4 1939.1 2044.3 1937.6 2045.8 C -1937 2046.1 1935.9 2046.1 1935.9 2046.8 C -1936.7 2046.3 1937.8 2046.2 1938.3 2045.6 C -f -S -n -1932.5 2040 m -1932.8 2038.1 1932 2038.9 1932.3 2040.3 C -1933.1 2040.3 1932.7 2041.7 1933.7 2041.5 C -1933.1 2041 1932.9 2040.5 1932.5 2040 C -f -S -n -2014.6 2035.2 m -2014.1 2033.6 2010.9 2034.7 2011.7 2032.4 C -2011.3 2031.9 2009.4 2030.7 2009.3 2032.1 C -2009.5 2033.7 2012.9 2033.8 2012.4 2035.7 C -2013 2036.4 2014.2 2036.5 2014.6 2035.2 C -f -S -n -1906.4 2030.7 m -1905 2031.6 1903.5 2033.6 1902 2032.8 C -1903.4 2034 1905.6 2031.4 1906.4 2030.7 C -f -S -n -1901.8 2037.2 m -1899.5 2034.8 1897.2 2032.5 1894.8 2030.2 C -1897.2 2032.5 1899.5 2034.8 1901.8 2037.2 C -[0.4 0.4 0 0] vc -f -S -n -1901.8 2032.4 m -1901.1 2031.6 1900.4 2030.7 1899.6 2030 C -1900.4 2030.7 1901.1 2031.6 1901.8 2032.4 C -[0 0 0 0.18] vc -f -S -n -1944.5 2030 m -1945.3 2029.9 1946.1 2029.8 1946.9 2029.7 C -1946.1 2029.8 1945.3 2029.9 1944.5 2030 C -f -S -n -vmrs -1997.8 2027.8 m -1997.7 2027.9 1997.6 2028.1 1997.3 2028 C -1997.4 2029.1 1998.5 2029.5 1999.2 2030 C -2000.1 2029.5 1998.9 2028 1997.8 2027.8 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1906.4 2029.2 m -1906.4 2026.6 1906.4 2024 1906.4 2021.3 C -1906.4 2024 1906.4 2026.6 1906.4 2029.2 C -f -S -n -2006.2 2025.9 m -2006 2025.9 2005.8 2025.8 2005.7 2025.6 C -2005.7 2025.5 2005.7 2025.3 2005.7 2025.2 C -2004.6 2025.8 2002.7 2024.7 2001.9 2026.1 C -2001.9 2027.9 2007.8 2029.2 2006.2 2025.9 C -[0 0 0 0] vc -f -S -n -1952.4 2026.8 m -1950.9 2027 1949.6 2026.4 1948.6 2027.3 C -1949.7 2026.1 1952.2 2027.7 1952.4 2026.8 C -[0 0 0 0.18] vc -f -S -n -1896.5 2026.8 m -1894.7 2025.1 1892.9 2023.3 1891 2021.6 C -1892.9 2023.3 1894.7 2025.1 1896.5 2026.8 C -f -S -n -1958.4 2024 m -1958.5 2025.5 1959.4 2024.8 1959.4 2026.4 C -1959.3 2025.3 1959.8 2024.1 1958.4 2024 C -f -S -n -1903.5 2019.2 m -1902.6 2018.6 1902.6 2020 1901.6 2019.9 C -1902.5 2020.8 1901.7 2021.4 1902.8 2022 C -1904.1 2022.2 1903.2 2020.1 1903.5 2019.2 C -f -S -n -1933 2018.7 m -1931.7 2019.6 1930.1 2021.6 1928.7 2020.8 C -1930.1 2022.1 1931.8 2019.5 1933 2018.7 C -f -S -n -1888.4 2016.3 m -1887.3 2018.4 1889.1 2019.6 1890.3 2020.8 C -1889.3 2019.5 1887.8 2018.3 1888.4 2016.3 C -f -S -n -1928.4 2020.4 m -1927.7 2019.6 1927 2018.7 1926.3 2018 C -1927 2018.7 1927.7 2019.6 1928.4 2020.4 C -f -S -n -vmrs -1911.2 2018.2 m -1909.8 2018.3 1907.3 2017.2 1906.8 2018.9 C -1908.1 2018.1 1910.5 2018.6 1911.2 2018.2 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1915.5 2015.1 m -1915.4 2013.9 1914 2013.3 1913.1 2012.9 C -1913.6 2013.9 1916.3 2013.8 1914.8 2015.3 C -1914.2 2015.6 1913.1 2015.6 1913.1 2016.3 C -1913.9 2015.9 1915 2015.7 1915.5 2015.1 C -f -S -n -1923.2 2014.8 m -1921.3 2013.1 1919.5 2011.3 1917.6 2009.6 C -1919.5 2011.3 1921.3 2013.1 1923.2 2014.8 C -f -S -n -1933 2012.7 m -1933 2011.7 1933 2010.8 1933 2009.8 C -1933 2010.8 1933 2011.7 1933 2012.7 C -f -S -n -1909.7 2008.1 m -1908.9 2009.2 1910.1 2009.9 1910.4 2011 C -1911.1 2010.7 1908.9 2009.7 1909.7 2008.1 C -f -S -n -1930.1 2007.2 m -1929.2 2006.6 1929.2 2008 1928.2 2007.9 C -1929.1 2008.8 1928.4 2009.4 1929.4 2010 C -1930.7 2010.2 1929.9 2008.1 1930.1 2007.2 C -f -S -n -1915 2004.3 m -1914 2006.4 1915.7 2007.6 1916.9 2008.8 C -1915.9 2007.5 1914.4 2006.3 1915 2004.3 C -f -S -n -1937.8 2006.2 m -1936.4 2006.3 1934 2005.2 1933.5 2006.9 C -1934.7 2006.1 1937.1 2006.6 1937.8 2006.2 C -f -S -n -1983.9 2006 m -1983.3 2004.3 1980.2 2005.4 1981 2003.1 C -1980.6 2002.7 1978.7 2001.5 1978.6 2002.8 C -1978.8 2004.4 1982.1 2004.5 1981.7 2006.4 C -1982.3 2007.2 1983.5 2007.2 1983.9 2006 C -f -S -n -1942.1 2003.1 m -1942 2001.9 1940.6 2001.3 1939.7 2000.9 C -1940.2 2001.9 1943 2001.8 1941.4 2003.3 C -1940.9 2003.6 1939.7 2003.6 1939.7 2004.3 C -1940.5 2003.9 1941.6 2003.7 1942.1 2003.1 C -f -S -n -vmrs -1967.1 1998.5 m -1967 1998.6 1966.8 1998.8 1966.6 1998.8 C -1966.7 1999.8 1967.8 2000.2 1968.5 2000.7 C -1969.4 2000.2 1968.2 1998.8 1967.1 1998.5 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1936.4 1997.6 m -1936.7 1995.6 1935.8 1996.4 1936.1 1997.8 C -1936.9 1997.9 1936.5 1999.2 1937.6 1999 C -1937 1998.5 1936.8 1998 1936.4 1997.6 C -f -S -n -1975.5 1996.6 m -1975.2 1996.7 1975.1 1996.5 1975 1996.4 C -1975 1996.2 1975 1996.1 1975 1995.9 C -1973.9 1996.5 1972 1995.5 1971.2 1996.8 C -1971.2 1998.6 1977 1999.9 1975.5 1996.6 C -[0 0 0 0] vc -f -S -n -1949.3 2097.4 m -1950.3 2096.9 1951.2 2096.4 1952.2 2096 C -1951.2 2096.4 1950.3 2096.9 1949.3 2097.4 C -[0.4 0.4 0 0] vc -f -S -n -1960.8 2091.6 m -1961.7 2091.2 1962.6 2090.9 1963.5 2090.4 C -1962.6 2090.9 1961.7 2091.2 1960.8 2091.6 C -f -S -n -1964.4 2090 m -1965.7 2089.2 1967 2088.5 1968.3 2087.8 C -1967 2088.5 1965.7 2089.2 1964.4 2090 C -f -S -n -1976 2083.7 m -1976.3 2082.3 1975.2 2079.1 1976.9 2079.4 C -1978.8 2080.7 1980.3 2082.9 1982.2 2084.2 C -1980.6 2083.1 1978.2 2080.2 1976 2078.9 C -1975.6 2081.2 1977 2084.9 1973.8 2085.4 C -1972.2 2086.1 1970.7 2087 1969 2087.6 C -1971.4 2086.5 1974.1 2085.6 1976 2083.7 C -f -S -n -1983.9 2084.2 m -1984.8 2083.7 1985.8 2083.2 1986.8 2082.8 C -1985.8 2083.2 1984.8 2083.7 1983.9 2084.2 C -f -S -n -1995.4 2078.4 m -1996.3 2078 1997.1 2077.7 1998 2077.2 C -1997.1 2077.7 1996.3 2078 1995.4 2078.4 C -f -S -n -1999 2076.8 m -2000.3 2076 2001.6 2075.3 2002.8 2074.6 C -2001.6 2075.3 2000.3 2076 1999 2076.8 C -f -S -n -vmrs -1929.6 2065.7 m -1930.1 2065.6 1929.8 2068.6 1929.9 2070 C -1929.8 2068.6 1930.1 2067 1929.6 2065.7 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1906.6 2049.4 m -1906.6 2046.7 1906.6 2043.9 1906.6 2041.2 C -1906.6 2043.9 1906.6 2046.7 1906.6 2049.4 C -f -S -n -2016 2047.5 m -2014.8 2048 2013.5 2048.3 2012.4 2049.4 C -2013.5 2048.3 2014.8 2048 2016 2047.5 C -f -S -n -2016.5 2047.2 m -2017.3 2046.9 2018.1 2046.6 2018.9 2046.3 C -2018.1 2046.6 2017.3 2046.9 2016.5 2047.2 C -f -S -n -1912.4 2028.5 m -1911.8 2032.4 1912.4 2037.2 1911.9 2041.2 C -1911.5 2037.2 1911.7 2032.9 1911.6 2028.8 C -1911.6 2033.5 1911.6 2038.9 1911.6 2042.9 C -1912.5 2042.2 1911.6 2043.9 1912.6 2043.6 C -1912.9 2039.3 1913.1 2033.3 1912.4 2028.5 C -[0.21 0.21 0 0] vc -f -S -n -1906.8 2040.8 m -1906.8 2039 1906.8 2037.2 1906.8 2035.5 C -1906.8 2037.2 1906.8 2039 1906.8 2040.8 C -[0.4 0.4 0 0] vc -f -S -n -1905.9 2035.2 m -1904.9 2036.4 1903.7 2037.2 1902.3 2037.4 C -1903.7 2037.2 1904.9 2036.4 1905.9 2035.2 C -f -S -n -1906.1 2031.2 m -1907 2031.1 1906.4 2028 1906.6 2030.7 C -1905.5 2032.1 1904 2032.8 1902.5 2033.6 C -1903.9 2033.2 1905 2032.1 1906.1 2031.2 C -f -S -n -1908.3 2018.7 m -1905.2 2018.6 1907.1 2023.2 1906.6 2025.4 C -1906.8 2023 1905.9 2019.5 1908.3 2018.7 C -f -S -n -1889.6 1998 m -1889 2001.9 1889.6 2006.7 1889.1 2010.8 C -1888.7 2006.7 1888.9 2002.4 1888.8 1998.3 C -1888.8 2003 1888.8 2008.4 1888.8 2012.4 C -1889.7 2011.7 1888.8 2013.4 1889.8 2013.2 C -1890.1 2008.8 1890.3 2002.8 1889.6 1998 C -[0.21 0.21 0 0] vc -f -S -n -vmrs -1999 2001.4 m -2001 2000.3 2003 1999.2 2005 1998 C -2003 1999.2 2001 2000.3 1999 2001.4 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1916.2 1986 m -1915.7 1989.9 1916.3 1994.7 1915.7 1998.8 C -1915.3 1994.7 1915.5 1990.4 1915.5 1986.3 C -1915.5 1991 1915.5 1996.4 1915.5 2000.4 C -1916.3 1999.7 1915.5 2001.4 1916.4 2001.2 C -1916.7 1996.8 1917 1990.8 1916.2 1986 C -[0.21 0.21 0 0] vc -f -S -n -1886.9 1989.6 m -1887.8 1989.2 1888.7 1988.9 1889.6 1988.4 C -1888.7 1988.9 1887.8 1989.2 1886.9 1989.6 C -[0.4 0.4 0 0] vc -f -S -n -1892.4 1986.8 m -1895.1 1985.1 1897.9 1983.6 1900.6 1982 C -1897.9 1983.6 1895.1 1985.1 1892.4 1986.8 C -f -S -n -1907.3 1979.3 m -1908.5 1978.9 1909.7 1978.5 1910.9 1978.1 C -1909.7 1978.5 1908.5 1978.9 1907.3 1979.3 C -f -S -n -1938.5 1966.6 m -1942.6 1970.1 1945.9 1976.4 1951.7 1979.1 C -1946.2 1976.1 1943.1 1970.9 1938.5 1966.6 C -f -S -n -1955.1 1978.6 m -1955.9 1978.2 1956.7 1977.8 1957.5 1977.4 C -1956.7 1977.8 1955.9 1978.2 1955.1 1978.6 C -f -S -n -1913.6 1977.6 m -1914.5 1977.2 1915.3 1976.9 1916.2 1976.4 C -1915.3 1976.9 1914.5 1977.2 1913.6 1977.6 C -f -S -n -1919.1 1974.8 m -1921.8 1973.1 1924.5 1971.6 1927.2 1970 C -1924.5 1971.6 1921.8 1973.1 1919.1 1974.8 C -f -S -n -1963.5 1974.5 m -1964.5 1974 1965.6 1973.4 1966.6 1972.8 C -1965.6 1973.4 1964.5 1974 1963.5 1974.5 C -f -S -n -vmrs -1967.8 1972.4 m -1970 1971.2 1972.1 1970 1974.3 1968.8 C -1972.1 1970 1970 1971.2 1967.8 1972.4 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1934 1967.3 m -1935.2 1966.9 1936.4 1966.5 1937.6 1966.1 C -1936.4 1966.5 1935.2 1966.9 1934 1967.3 C -f -S -n -1928.9 2061.2 m -1928.9 2059.2 1928.9 2057.3 1928.9 2055.4 C -1928.9 2057.3 1928.9 2059.2 1928.9 2061.2 C -[0.21 0.21 0 0] vc -f -S -n -1917.2 2047 m -1917.8 2046.5 1919.6 2046.8 1920 2047.2 C -1920 2046.5 1920.9 2046.8 1921 2046.3 C -1921.9 2047.3 1921.3 2044.1 1921.5 2044.1 C -1919.7 2044.8 1915.7 2043.5 1916.2 2046 C -1916.2 2048.3 1917 2045.9 1917.2 2047 C -[0 0 0 0] vc -f -S -n -1922 2044.1 m -1923.5 2043.2 1927 2045.4 1927.5 2042.9 C -1927.1 2042.6 1927.3 2040.9 1927.2 2041.5 C -1924.9 2042.3 1920.9 2040.6 1922 2044.1 C -f -S -n -1934.9 2043.9 m -1935.2 2043.4 1934.4 2042.7 1934 2042.2 C -1933.2 2041.8 1932.4 2042.8 1932.8 2043.2 C -1932.9 2044 1934.3 2043.3 1934.9 2043.9 C -f -S -n -1906.1 2030.7 m -1906.1 2028.8 1906.1 2027 1906.1 2025.2 C -1906.1 2027 1906.1 2028.8 1906.1 2030.7 C -[0.21 0.21 0 0] vc -f -S -n -1932.8 2018.7 m -1932.8 2016.8 1932.8 2014.8 1932.8 2012.9 C -1932.8 2014.8 1932.8 2016.8 1932.8 2018.7 C -f -S -n -1894.4 2016.5 m -1895 2016 1896.8 2016.3 1897.2 2016.8 C -1897.2 2016 1898.1 2016.3 1898.2 2015.8 C -1899.1 2016.8 1898.5 2013.6 1898.7 2013.6 C -1896.9 2014.4 1892.9 2013 1893.4 2015.6 C -1893.4 2017.8 1894.2 2015.4 1894.4 2016.5 C -[0 0 0 0] vc -f -S -n -1899.2 2013.6 m -1900.7 2012.7 1904.2 2014.9 1904.7 2012.4 C -1904.3 2012.1 1904.5 2010.5 1904.4 2011 C -1902.1 2011.8 1898.1 2010.1 1899.2 2013.6 C -f -S -n -vmrs -1912.1 2013.4 m -1912.4 2012.9 1911.6 2012.3 1911.2 2011.7 C -1910.4 2011.4 1909.6 2012.3 1910 2012.7 C -1910.1 2013.5 1911.5 2012.9 1912.1 2013.4 C -[0 0 0 0] vc -f -0.4 w -2 J -2 M -S -n -1921 2004.5 m -1921.6 2004 1923.4 2004.3 1923.9 2004.8 C -1923.8 2004 1924.8 2004.3 1924.8 2003.8 C -1925.7 2004.8 1925.1 2001.6 1925.3 2001.6 C -1923.6 2002.4 1919.6 2001 1920 2003.6 C -1920 2005.8 1920.8 2003.4 1921 2004.5 C -f -S -n -1925.8 2001.6 m -1927.3 2000.7 1930.8 2002.9 1931.3 2000.4 C -1930.9 2000.1 1931.1 1998.5 1931.1 1999 C -1928.7 1999.8 1924.8 1998.1 1925.8 2001.6 C -f -S -n -1938.8 2001.4 m -1939 2000.9 1938.2 2000.3 1937.8 1999.7 C -1937.1 1999.4 1936.2 2000.3 1936.6 2000.7 C -1936.7 2001.5 1938.1 2000.9 1938.8 2001.4 C -f -S -n -1908.6691 2008.1348 m -1897.82 2010.0477 L -1894.1735 1989.3671 L -1905.0226 1987.4542 L -1908.6691 2008.1348 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1895.041763 1994.291153 m -0 0 32 0 0 (l) ts -} -true -[0 0 0 1]sts -Q -1979.2185 1991.7809 m -1960.6353 1998.5452 L -1953.4532 1978.8124 L -1972.0363 1972.0481 L -1979.2185 1991.7809 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [18.793335 -6.84082 6.84021 18.793335 0 0] makesetfont -1955.163254 1983.510773 m -0 0 32 0 0 (\256) ts -} -true -[0 0 0 1]sts -Q -1952.1544 2066.5423 m -1938.0739 2069.025 L -1934.4274 2048.3444 L -1948.5079 2045.8617 L -1952.1544 2066.5423 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1935.29567 2053.268433 m -0 0 32 0 0 (") ts -} -true -[0 0 0 1]sts -Q -1931.7231 2043.621 m -1919.3084 2048.14 L -1910.6898 2024.4607 L -1923.1046 2019.9417 L -1931.7231 2043.621 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [22.552002 -8.208984 8.208252 22.552002 0 0] makesetfont -1912.741867 2030.098648 m -0 0 32 0 0 (=) ts -} -true -[0 0 0 1]sts -Q -1944 2024.5 m -1944 2014 L -0.8504 w -0 J -3.863693 M -[0 0 0 1] vc -false setoverprint -S -n -1944.25 2019.1673 m -1952.5 2015.9173 L -S -n -1931.0787 2124.423 m -1855.5505 2043.4285 L -1871.0419 2013.0337 L -1946.5701 2094.0282 L -1931.0787 2124.423 L -n -q -_bfh -%%IncludeResource: font ZapfHumanist601BT-Bold -_efh -{ -f1 [22.155762 23.759277 -14.753906 28.947754 0 0] makesetfont -1867.35347 2020.27063 m -0 0 32 0 0 (Isabelle) ts -} -true -[0 0 0 1]sts -Q -1933.5503 1996.9547 m -1922.7012 1998.8677 L -1919.0547 1978.1871 L -1929.9038 1976.2741 L -1933.5503 1996.9547 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1919.922913 1983.111069 m -0 0 32 0 0 (b) ts -} -true -[0 0 0 1]sts -Q -2006.3221 2025.7184 m -1993.8573 2027.9162 L -1990.2108 2007.2356 L -2002.6756 2005.0378 L -2006.3221 2025.7184 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1991.07901 2012.159653 m -0 0 32 0 0 (a) ts -} -true -[0 0 0 1]sts -Q -vmrs -2030.0624 2094.056 m -1956.3187 2120.904 L -1956.321 2095.3175 L -2030.0647 2068.4695 L -2030.0624 2094.056 L -n -q -_bfh -%%IncludeResource: font ZapfHumanist601BT-Bold -_efh -{ -f1 [22.898804 -8.336792 -0.002197 24.368408 0 0] makesetfont -1956.320496 2101.409561 m -0 0 32 0 0 (Isar) ts -} -true -[0 0 0 1]sts -Q -vmr -vmr -end -%%Trailer -%%DocumentNeededResources: font Symbol -%%+ font ZapfHumanist601BT-Bold -%%DocumentFonts: Symbol -%%+ ZapfHumanist601BT-Bold -%%DocumentNeededFonts: Symbol -%%+ ZapfHumanist601BT-Bold diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/isabelle_isar.pdf Binary file doc-src/IsarAdvanced/Functions/isabelle_isar.pdf has changed diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/mathpartir.sty --- a/doc-src/IsarAdvanced/Functions/mathpartir.sty Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,421 +0,0 @@ -% Mathpartir --- Math Paragraph for Typesetting Inference Rules -% -% Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy -% -% Author : Didier Remy -% Version : 1.2.0 -% Bug Reports : to author -% Web Site : http://pauillac.inria.fr/~remy/latex/ -% -% Mathpartir is free software; you can redistribute it and/or modify -% it under the terms of the GNU General Public License as published by -% the Free Software Foundation; either version 2, or (at your option) -% any later version. -% -% Mathpartir is distributed in the hope that it will be useful, -% but WITHOUT ANY WARRANTY; without even the implied warranty of -% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the -% GNU General Public License for more details -% (http://pauillac.inria.fr/~remy/license/GPL). -% -%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% -% File mathpartir.sty (LaTeX macros) -%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% - -\NeedsTeXFormat{LaTeX2e} -\ProvidesPackage{mathpartir} - [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules] - -%% - -%% Identification -%% Preliminary declarations - -\RequirePackage {keyval} - -%% Options -%% More declarations - -%% PART I: Typesetting maths in paragraphe mode - -\newdimen \mpr@tmpdim - -% To ensure hevea \hva compatibility, \hva should expands to nothing -% in mathpar or in inferrule -\let \mpr@hva \empty - -%% normal paragraph parametters, should rather be taken dynamically -\def \mpr@savepar {% - \edef \MathparNormalpar - {\noexpand \lineskiplimit \the\lineskiplimit - \noexpand \lineskip \the\lineskip}% - } - -\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em} -\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em} -\def \mpr@lineskip {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em} -\let \MathparLineskip \mpr@lineskip -\def \mpr@paroptions {\MathparLineskip} -\let \mpr@prebindings \relax - -\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em - -\def \mpr@goodbreakand - {\hskip -\mpr@andskip \penalty -1000\hskip \mpr@andskip} -\def \mpr@and {\hskip \mpr@andskip} -\def \mpr@andcr {\penalty 50\mpr@and} -\def \mpr@cr {\penalty -10000\mpr@and} -\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10} - -\def \mpr@bindings {% - \let \and \mpr@andcr - \let \par \mpr@andcr - \let \\\mpr@cr - \let \eqno \mpr@eqno - \let \hva \mpr@hva - } -\let \MathparBindings \mpr@bindings - -% \@ifundefined {ignorespacesafterend} -% {\def \ignorespacesafterend {\aftergroup \ignorespaces} - -\newenvironment{mathpar}[1][] - {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering - \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else - \noindent $\displaystyle\fi - \MathparBindings} - {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend} - -% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum -% \wd0 < \hsize $$\box0$$\else \bmathpar #1\emathpar \fi} - -%%% HOV BOXES - -\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip - \vbox \bgroup \tabskip 0em \let \\ \cr - \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup - \egroup} - -\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize - \box0\else \mathvbox {#1}\fi} - - -%% Part II -- operations on lists - -\newtoks \mpr@lista -\newtoks \mpr@listb - -\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter -{#2}\edef #2{\the \mpr@lista \the \mpr@listb}} - -\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter -{#2}\edef #2{\the \mpr@listb\the\mpr@lista}} - -\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb -\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}} - -\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2} -\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}} - -\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2} -\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}} - -\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}% - \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the - \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty - \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop - \mpr@flatten \mpr@all \mpr@to \mpr@one - \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof - \mpr@all \mpr@stripend - \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi - \ifx 1\mpr@isempty - \repeat -} - -\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty - \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp} - -%% Part III -- Type inference rules - -\newif \if@premisse -\newbox \mpr@hlist -\newbox \mpr@vlist -\newif \ifmpr@center \mpr@centertrue -\def \mpr@htovlist {% - \setbox \mpr@hlist - \hbox {\strut - \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi - \unhbox \mpr@hlist}% - \setbox \mpr@vlist - \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist - \else \unvbox \mpr@vlist \box \mpr@hlist - \fi}% -} -% OLD version -% \def \mpr@htovlist {% -% \setbox \mpr@hlist -% \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}% -% \setbox \mpr@vlist -% \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist -% \else \unvbox \mpr@vlist \box \mpr@hlist -% \fi}% -% } - -\def \mpr@item #1{$\displaystyle #1$} -\def \mpr@sep{2em} -\def \mpr@blank { } -\def \mpr@hovbox #1#2{\hbox - \bgroup - \ifx #1T\@premissetrue - \else \ifx #1B\@premissefalse - \else - \PackageError{mathpartir} - {Premisse orientation should either be T or B} - {Fatal error in Package}% - \fi \fi - \def \@test {#2}\ifx \@test \mpr@blank\else - \setbox \mpr@hlist \hbox {}% - \setbox \mpr@vlist \vbox {}% - \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi - \let \@hvlist \empty \let \@rev \empty - \mpr@tmpdim 0em - \expandafter \mpr@makelist #2\mpr@to \mpr@flat - \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi - \def \\##1{% - \def \@test {##1}\ifx \@test \empty - \mpr@htovlist - \mpr@tmpdim 0em %%% last bug fix not extensively checked - \else - \setbox0 \hbox{\mpr@item {##1}}\relax - \advance \mpr@tmpdim by \wd0 - %\mpr@tmpdim 1.02\mpr@tmpdim - \ifnum \mpr@tmpdim < \hsize - \ifnum \wd\mpr@hlist > 0 - \if@premisse - \setbox \mpr@hlist - \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}% - \else - \setbox \mpr@hlist - \hbox {\unhbox \mpr@hlist \hskip \mpr@sep \unhbox0}% - \fi - \else - \setbox \mpr@hlist \hbox {\unhbox0}% - \fi - \else - \ifnum \wd \mpr@hlist > 0 - \mpr@htovlist - \mpr@tmpdim \wd0 - \fi - \setbox \mpr@hlist \hbox {\unhbox0}% - \fi - \advance \mpr@tmpdim by \mpr@sep - \fi - }% - \@rev - \mpr@htovlist - \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist - \fi - \egroup -} - -%%% INFERENCE RULES - -\@ifundefined{@@over}{% - \let\@@over\over % fallback if amsmath is not loaded - \let\@@overwithdelims\overwithdelims - \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims - \let\@@above\above \let\@@abovewithdelims\abovewithdelims - }{} - -%% The default - -\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em - $\displaystyle {#1\mpr@over #2}$}} -\let \mpr@fraction \mpr@@fraction - -%% A generic solution to arrow - -\def \mpr@make@fraction #1#2#3#4#5{\hbox {% - \def \mpr@tail{#1}% - \def \mpr@body{#2}% - \def \mpr@head{#3}% - \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}% - \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}% - \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}% - \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax - \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax - \setbox0=\hbox {$\box1 \@@atop \box2$}% - \dimen0=\wd0\box0 - \box0 \hskip -\dimen0\relax - \hbox to \dimen0 {$% - \mathrel{\mpr@tail}\joinrel - \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}% - $}}} - -%% Old stuff should be removed in next version -\def \mpr@@reduce #1#2{\hbox - {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}} -\def \mpr@@rewrite #1#2#3{\hbox - {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}} -\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}} - -\def \mpr@empty {} -\def \mpr@inferrule - {\bgroup - \ifnum \linewidth<\hsize \hsize \linewidth\fi - \mpr@rulelineskip - \let \and \qquad - \let \hva \mpr@hva - \let \@rulename \mpr@empty - \let \@rule@options \mpr@empty - \let \mpr@over \@@over - \mpr@inferrule@} -\newcommand {\mpr@inferrule@}[3][] - {\everymath={\displaystyle}% - \def \@test {#2}\ifx \empty \@test - \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}% - \else - \def \@test {#3}\ifx \empty \@test - \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}% - \else - \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}% - \fi \fi - \def \@test {#1}\ifx \@test\empty \box0 - \else \vbox -%%% Suggestion de Francois pour les etiquettes longues -%%% {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi - {\hbox {\RefTirName {#1}}\box0}\fi - \egroup} - -\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}} - -% They are two forms -% \inferrule [label]{[premisses}{conclusions} -% or -% \inferrule* [options]{[premisses}{conclusions} -% -% Premisses and conclusions are lists of elements separated by \\ -% Each \\ produces a break, attempting horizontal breaks if possible, -% and vertical breaks if needed. -% -% An empty element obtained by \\\\ produces a vertical break in all cases. -% -% The former rule is aligned on the fraction bar. -% The optional label appears on top of the rule -% The second form to be used in a derivation tree is aligned on the last -% line of its conclusion -% -% The second form can be parameterized, using the key=val interface. The -% folloiwng keys are recognized: -% -% width set the width of the rule to val -% narrower set the width of the rule to val\hsize -% before execute val at the beginning/left -% lab put a label [Val] on top of the rule -% lskip add negative skip on the right -% left put a left label [Val] -% Left put a left label [Val], ignoring its width -% right put a right label [Val] -% Right put a right label [Val], ignoring its width -% leftskip skip negative space on the left-hand side -% rightskip skip negative space on the right-hand side -% vdots lift the rule by val and fill vertical space with dots -% after execute val at the end/right -% -% Note that most options must come in this order to avoid strange -% typesetting (in particular leftskip must preceed left and Left and -% rightskip must follow Right or right; vdots must come last -% or be only followed by rightskip. -% - -%% Keys that make sence in all kinds of rules -\def \mprset #1{\setkeys{mprset}{#1}} -\define@key {mprset}{flushleft}[]{\mpr@centerfalse} -\define@key {mprset}{center}[]{\mpr@centertrue} -\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} -\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1} -\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} - -\newbox \mpr@right -\define@key {mpr}{flushleft}[]{\mpr@centerfalse} -\define@key {mpr}{center}[]{\mpr@centertrue} -\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} -\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1} -\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} -\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax - \advance \hsize by -\wd0\box0} -\define@key {mpr}{width}{\hsize #1} -\define@key {mpr}{sep}{\def\mpr@sep{#1}} -\define@key {mpr}{before}{#1} -\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} -\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} -\define@key {mpr}{narrower}{\hsize #1\hsize} -\define@key {mpr}{leftskip}{\hskip -#1} -\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce} -\define@key {mpr}{rightskip} - {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}} -\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax - \advance \hsize by -\wd0\box0} -\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax - \advance \hsize by -\wd0\box0} -\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}} -\define@key {mpr}{right} - {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0 - \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} -\define@key {mpr}{RIGHT} - {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0 - \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} -\define@key {mpr}{Right} - {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}} -\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}} -\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}} - -\newdimen \rule@dimen -\newcommand \mpr@inferstar@ [3][]{\setbox0 - \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax - \setbox \mpr@right \hbox{}% - $\setkeys{mpr}{#1}% - \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else - \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi - \box \mpr@right \mpr@vdots$} - \setbox1 \hbox {\strut} - \rule@dimen \dp0 \advance \rule@dimen by -\dp1 - \raise \rule@dimen \box0} - -\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}} -\newcommand \mpr@err@skipargs[3][]{} -\def \mpr@inferstar*{\ifmmode - \let \@do \mpr@inferstar@ - \else - \let \@do \mpr@err@skipargs - \PackageError {mathpartir} - {\string\inferrule* can only be used in math mode}{}% - \fi \@do} - - -%%% Exports - -% Envirnonment mathpar - -\let \inferrule \mpr@infer - -% make a short name \infer is not already defined -\@ifundefined {infer}{\let \infer \mpr@infer}{} - -\def \TirNameStyle #1{\small \textsc{#1}} -\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}} -\let \TirName \tir@name -\let \DefTirName \TirName -\let \RefTirName \TirName - -%%% Other Exports - -% \let \listcons \mpr@cons -% \let \listsnoc \mpr@snoc -% \let \listhead \mpr@head -% \let \listmake \mpr@makelist - - - - -\endinput diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Functions/style.sty --- a/doc-src/IsarAdvanced/Functions/style.sty Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,46 +0,0 @@ -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} -\newcommand{\chref}[1]{chapter~\ref{#1}} -\newcommand{\figref}[1]{figure~\ref{#1}} - -%% math -\newcommand{\text}[1]{\mbox{#1}} -\newcommand{\isasymvartheta}{\isamath{\theta}} -\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}} - -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\sloppy -\binperiod -\underscoreon - -\renewcommand{\isadigit}[1]{\isamath{#1}} - -\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} - -\isafoldtag{FIXME} -\isakeeptag{mlref} -\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} -\renewcommand{\endisatagmlref}{\endgroup} - -\newcommand{\isasymGUESS}{\isakeyword{guess}} -\newcommand{\isasymOBTAIN}{\isakeyword{obtain}} -\newcommand{\isasymTHEORY}{\isakeyword{theory}} -\newcommand{\isasymUSES}{\isakeyword{uses}} -\newcommand{\isasymEND}{\isakeyword{end}} -\newcommand{\isasymCONSTS}{\isakeyword{consts}} -\newcommand{\isasymDEFS}{\isakeyword{defs}} -\newcommand{\isasymTHEOREM}{\isakeyword{theorem}} -\newcommand{\isasymDEFINITION}{\isakeyword{definition}} - -\isabellestyle{it} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 2775062fd3a9 -r 2f4684e2ea95 doc-src/IsarAdvanced/Makefile.in --- a/doc-src/IsarAdvanced/Makefile.in Mon Mar 02 16:58:39 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,13 +0,0 @@ -# $Id$ -# - -include ../../Makefile.in - -SEDINDEX = ../../sedindex -FIXBOOKMARKS = perl -pi ../../fixbookmarks.pl - -isabelle_isar.eps: - test -r isabelle_isar.eps || ln -s ../../gfx/isabelle_isar.eps . - -isabelle_isar.pdf: - test -r isabelle_isar.pdf || ln -s ../../gfx/isabelle_isar.pdf .